src/Tools/isac/Knowledge/Rational.thy
author Walther Neuper <neuper@ist.tugraz.at>
Sat, 24 Aug 2013 11:54:39 +0200
changeset 52087 dacbaaea9f95
parent 52075 97db4e3d07ae
child 52089 7a740eedefc0
permissions -rwxr-xr-x
GCD_Poly_ML: conversions term <--> poly

not contained in GCD_Poly_ML, because the latter is meant for TUM,
while conversions use lots of Isac tools.
neuper@37906
     1
(* rationals, i.e. fractions of multivariate polynomials over the real field
neuper@37906
     2
   author: isac team
neuper@37906
     3
   Copyright (c) isac team 2002
neuper@37906
     4
   Use is subject to license terms.
neuper@37906
     5
neuper@37906
     6
   depends on Poly (and not on Atools), because 
neuper@37906
     7
   fractions with _normalized_ polynomials are canceled, added, etc.
neuper@48880
     8
neuper@48880
     9
   ATTENTION WN130616: WITH Unsynchronized.ref Rational.thy TAKES ~1MIN FOR EVALUATION
neuper@37906
    10
*)
neuper@37906
    11
neuper@48884
    12
theory Rational 
neuper@52075
    13
imports Poly "~~/src/Tools/isac/Knowledge/GCD_Poly_ML"
neuper@48884
    14
begin
neuper@37906
    15
neuper@37906
    16
consts
neuper@37906
    17
jan@42300
    18
  is'_expanded    :: "real => bool" ("_ is'_expanded")     (*RL->Poly.thy*)
jan@42300
    19
  is'_ratpolyexp  :: "real => bool" ("_ is'_ratpolyexp") 
jan@42300
    20
  get_denominator :: "real => real"
jan@42338
    21
  get_numerator   :: "real => real"
jan@42344
    22
  
neuper@37906
    23
t@42211
    24
axioms(*axiomatization where*) (*.not contained in Isabelle2002,
neuper@37950
    25
          stated as axioms, TODO?: prove as theorems*)
neuper@37950
    26
t@42211
    27
  mult_cross:      "[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)" (*and*)
t@42211
    28
  mult_cross1:     "   b ~= 0            ==> (a / b = c    ) = (a     = b * c)" (*and*)
t@42211
    29
  mult_cross2:     "           d ~= 0    ==> (a     = c / d) = (a * d =     c)" (*and*)
neuper@37950
    30
                  
t@42211
    31
  add_minus:       "a + b - b = a"(*RL->Poly.thy*) (*and*)
t@42211
    32
  add_minus1:      "a - b + b = a"(*RL->Poly.thy*) (*and*)
neuper@37950
    33
                  
t@42211
    34
  rat_mult:        "a / b * (c / d) = a * c / (b * d)"(*?Isa02*)  (*and*)
t@42211
    35
  rat_mult2:       "a / b *  c      = a * c /  b     "(*?Isa02*) (*and*)
t@42211
    36
t@42211
    37
  rat_mult_poly_l: "c is_polyexp ==> c * (a / b) = c * a /  b" (*and*)
t@42211
    38
  rat_mult_poly_r: "c is_polyexp ==> (a / b) * c = a * c /  b" (*and*)
neuper@37906
    39
neuper@37906
    40
(*real_times_divide1_eq .. Isa02*) 
t@42211
    41
  real_times_divide_1_eq:  "-1    * (c / d) =-1 * c /      d " (*and*)
neuper@37978
    42
  real_times_divide_num:   "a is_const ==> 
t@42211
    43
	          	   a     * (c / d) = a * c /      d " (*and*)
t@42211
    44
t@42211
    45
  real_mult_div_cancel2:   "k ~= 0 ==> m * k / (n * k) = m / n" (*and*)
neuper@37978
    46
(*real_mult_div_cancel1:   "k ~= 0 ==> k * m / (k * n) = m / n"..Isa02*)
neuper@37906
    47
			  
t@42211
    48
  real_divide_divide1:     "y ~= 0 ==> (u / v) / (y / z) = (u / v) * (z / y)" (*and*)
t@42211
    49
  real_divide_divide1_mg:  "y ~= 0 ==> (u / v) / (y / z) = (u * z) / (y * v)" (*and*)
neuper@37978
    50
(*real_divide_divide2_eq:  "x / y / z = x / (y * z)"..Isa02*)
neuper@37906
    51
			  
t@42211
    52
  rat_power:               "(a / b)^^^n = (a^^^n) / (b^^^n)" (*and*)
neuper@37978
    53
neuper@37978
    54
neuper@37978
    55
  rat_add:         "[| a is_const; b is_const; c is_const; d is_const |] ==> 
t@42211
    56
	           a / c + b / d = (a * d + b * c) / (c * d)" (*and*)
neuper@37978
    57
  rat_add_assoc:   "[| a is_const; b is_const; c is_const; d is_const |] ==> 
t@42211
    58
	           a / c +(b / d + e) = (a * d + b * c)/(d * c) + e" (*and*)
neuper@37978
    59
  rat_add1:        "[| a is_const; b is_const; c is_const |] ==> 
t@42211
    60
	           a / c + b / c = (a + b) / c" (*and*)
neuper@37978
    61
  rat_add1_assoc:   "[| a is_const; b is_const; c is_const |] ==> 
t@42211
    62
	           a / c + (b / c + e) = (a + b) / c + e" (*and*)
neuper@37978
    63
  rat_add2:        "[| a is_const; b is_const; c is_const |] ==> 
t@42211
    64
	           a / c + b = (a + b * c) / c" (*and*)
neuper@37978
    65
  rat_add2_assoc:  "[| a is_const; b is_const; c is_const |] ==> 
t@42211
    66
	           a / c + (b + e) = (a + b * c) / c + e" (*and*)
neuper@37978
    67
  rat_add3:        "[| a is_const; b is_const; c is_const |] ==> 
t@42211
    68
	           a + b / c = (a * c + b) / c" (*and*)
neuper@37978
    69
  rat_add3_assoc:   "[| a is_const; b is_const; c is_const |] ==> 
neuper@37950
    70
	           a + (b / c + e) = (a * c + b) / c + e"
neuper@37950
    71
neuper@52087
    72
section {* Cancellation and addition of fractions *}
neuper@52087
    73
subsection {* Auxiliary functions and data *}
neuper@52087
    74
subsubsection {* Conversion term <--> poly *}
neuper@52087
    75
ML {*
neuper@52087
    76
fun monom_of_term  vs (c, es) (Free (id, _)) =
neuper@52087
    77
    if is_numeral id 
neuper@52087
    78
    then (id |> int_of_str |> the |> curry op * c, es) (*several numerals in one monom*)
neuper@52087
    79
    else (c, list_update es (find_index (curry op = id) vs) 1)
neuper@52087
    80
  | monom_of_term  vs (c, es) (Const ("Atools.pow", _) $ Free (id, _) $ Free (e, _)) =
neuper@52087
    81
    (c, list_update es (find_index (curry op = id) vs) (the (int_of_str e)))
neuper@52087
    82
  | monom_of_term vs (c, es) (Const ("Groups.times_class.times", _) $ m1 $ m2) =
neuper@52087
    83
    let val (c', es') = monom_of_term vs (c, es) m1
neuper@52087
    84
    in monom_of_term vs (c', es') m2 end
neuper@52087
    85
  | monom_of_term _ _ t = raise ERROR ("poly malformed with " ^ term2str t)
neuper@52087
    86
neuper@52087
    87
fun monoms_of_term vs (t as Free _) =
neuper@52087
    88
    [monom_of_term  vs (1, replicate (length vs) 0) t]
neuper@52087
    89
  | monoms_of_term vs (t as Const ("Atools.pow", _) $ _ $  _) =
neuper@52087
    90
    [monom_of_term  vs (1, replicate (length vs) 0) t]
neuper@52087
    91
  | monoms_of_term vs (t as Const ("Groups.times_class.times", _) $ _ $  _) =
neuper@52087
    92
    [monom_of_term  vs (1, replicate (length vs) 0) t]
neuper@52087
    93
  | monoms_of_term vs (Const ("Groups.plus_class.plus", _) $ ms1 $ ms2) =
neuper@52087
    94
    (monoms_of_term vs ms1) @ (monoms_of_term vs ms2)
neuper@52087
    95
  | monoms_of_term _ t = raise ERROR ("poly malformed with " ^ term2str t)
neuper@52087
    96
neuper@52087
    97
(* convert a term to the internal representation of a multivariate polynomial;
neuper@52087
    98
  the conversion is quite liberal, see test --- fun poly_of_term ---:
neuper@52087
    99
* the order of variables and the parentheses within a monomial are arbitrary
neuper@52087
   100
* the coefficient may be somewhere
neuper@52087
   101
* he order and the parentheses within monomials are arbitrary
neuper@52087
   102
But the term must be completely expand + over * (laws of distributivity are not applicable).
neuper@52087
   103
neuper@52087
   104
The function requires the free variables as strings already given, 
neuper@52087
   105
because the gcd involves 2 polynomials (with the same length for their list of exponents).
neuper@52087
   106
*)
neuper@52087
   107
fun poly_of_term vs (t as Const ("Groups.plus_class.plus", _) $ _ $ _) =
neuper@52087
   108
    (SOME (t |> monoms_of_term  vs |> order)
neuper@52087
   109
      handle ERROR _ => NONE)
neuper@52087
   110
  | poly_of_term vs t =
neuper@52087
   111
    (SOME [monom_of_term  vs (1, replicate (length vs) 0) t]
neuper@52087
   112
      handle ERROR _ => NONE)
neuper@52087
   113
neuper@52087
   114
fun is_poly t =
neuper@52087
   115
  let val vs = t |> vars |> map free2str |> sort string_ord
neuper@52087
   116
  in 
neuper@52087
   117
    case poly_of_term vs t of SOME _ => true | NONE => false
neuper@52087
   118
  end
neuper@52087
   119
val is_expanded = is_poly
neuper@52087
   120
neuper@52087
   121
(* convert internal representation of a multivariate polynomial to a term*)
neuper@52087
   122
fun term_of_es _ _ _ [] = [] (*assumes same length for vs and es*)
neuper@52087
   123
  | term_of_es baseT expT (_ :: vs) (0 :: es) =
neuper@52087
   124
    [] @ term_of_es baseT expT vs es
neuper@52087
   125
  | term_of_es baseT expT (v :: vs) (1 :: es) =
neuper@52087
   126
    [(Free (v, baseT))] @ term_of_es baseT expT vs es
neuper@52087
   127
  | term_of_es baseT expT (v :: vs) (e :: es) =
neuper@52087
   128
    [Const ("Atools.pow", [baseT, expT] ---> baseT) $ 
neuper@52087
   129
      (Free (v, baseT)) $  (Free (isastr_of_int e, expT))]
neuper@52087
   130
    @ term_of_es baseT expT vs es
neuper@52087
   131
neuper@52087
   132
fun term_of_monom baseT expT vs ((c, es): monom) =
neuper@52087
   133
    let val es' = term_of_es baseT expT vs es
neuper@52087
   134
    in 
neuper@52087
   135
      if c = 1 
neuper@52087
   136
      then 
neuper@52087
   137
        if es' = [] (*if es = [0,0,0,...]*)
neuper@52087
   138
        then Free (isastr_of_int c, baseT)
neuper@52087
   139
        else foldl (HOLogic.mk_binop "Groups.times_class.times") (hd es', tl es')
neuper@52087
   140
      else foldl (HOLogic.mk_binop "Groups.times_class.times") (Free (isastr_of_int c, baseT), es') 
neuper@52087
   141
    end
neuper@52087
   142
neuper@52087
   143
fun term_of_poly baseT expT vs p =
neuper@52087
   144
  let val monos = map (term_of_monom baseT expT vs) p
neuper@52087
   145
  in foldl (HOLogic.mk_binop "Groups.plus_class.plus") (hd monos, tl monos) end
neuper@52087
   146
*}
neuper@52087
   147
neuper@37950
   148
text {*calculate in rationals: gcd, lcm, etc.
neuper@37950
   149
      (c) Stefan Karnel 2002
neuper@37950
   150
      Institute for Mathematics D and Institute for Software Technology, 
neuper@37950
   151
      TU-Graz SS 2002 *}
neuper@37950
   152
neuper@37950
   153
text {* Remark on notions in the documentation below:
neuper@37950
   154
    referring to the remark on 'polynomials' in Poly.sml we use
neuper@37950
   155
    [2] 'polynomial' normalform (Polynom)
neuper@37950
   156
    [3] 'expanded_term' normalform (Ausmultiplizierter Term),
neuper@37950
   157
    where normalform [2] is a special case of [3], i.e. [3] implies [2].
neuper@37950
   158
    Instead of 
neuper@37950
   159
      'fraction with numerator and nominator both in normalform [2]'
neuper@37950
   160
      'fraction with numerator and nominator both in normalform [3]' 
neuper@37950
   161
    we say: 
neuper@37950
   162
      'fraction in normalform [2]'
neuper@37950
   163
      'fraction in normalform [3]' 
neuper@37950
   164
    or
neuper@37950
   165
      'fraction [2]'
neuper@37950
   166
      'fraction [3]'.
neuper@37950
   167
    a 'simple fraction' is a term with '/' as outmost operator and
neuper@37950
   168
    numerator and nominator in normalform [2] or [3].
neuper@37950
   169
*}
neuper@37950
   170
neuper@37979
   171
ML {* 
neuper@37972
   172
val thy = @{theory};
neuper@37972
   173
neuper@37950
   174
signature RATIONALI =
neuper@37950
   175
sig
neuper@37950
   176
  type mv_monom
neuper@37950
   177
  type mv_poly 
neuper@37950
   178
  val add_fraction_ : theory -> term -> (term * term list) option      
neuper@37950
   179
  val add_fraction_p_ : theory -> term -> (term * term list) option       
neuper@37950
   180
  val calculate_Rational : rls
neuper@37950
   181
  val calc_rat_erls:rls
neuper@37950
   182
  val cancel : rls
neuper@37950
   183
  val cancel_ : theory -> term -> (term * term list) option    
neuper@37950
   184
  val cancel_p : rls   
neuper@37950
   185
  val cancel_p_ : theory -> term -> (term * term list) option
neuper@37950
   186
  val common_nominator : rls              
neuper@37950
   187
  val common_nominator_ : theory -> term -> (term * term list) option
neuper@37950
   188
  val common_nominator_p : rls              
neuper@37950
   189
  val common_nominator_p_ : theory -> term -> (term * term list) option
neuper@37950
   190
  val eval_is_expanded : string -> 'a -> term -> theory -> 
neuper@37950
   191
			 (string * term) option                    
neuper@37950
   192
  val expanded2polynomial : term -> term option
neuper@37950
   193
  val factout_ : theory -> term -> (term * term list) option
neuper@37950
   194
  val factout_p_ : theory -> term -> (term * term list) option
neuper@37950
   195
  val is_expanded : term -> bool
neuper@37950
   196
  val is_polynomial : term -> bool
neuper@37950
   197
neuper@37950
   198
  val mv_gcd : (int * int list) list -> mv_poly -> mv_poly
neuper@37950
   199
  val mv_lcm : mv_poly -> mv_poly -> mv_poly
neuper@37950
   200
neuper@37950
   201
  val norm_expanded_rat_ : theory -> term -> (term * term list) option
neuper@37950
   202
(*WN0602.2.6.pull into struct !!!
neuper@37950
   203
  val norm_Rational : rls(*.normalizes an arbitrary rational term without
neuper@37950
   204
                           roots into a simple and canceled fraction
neuper@37950
   205
                           with normalform [2].*)
neuper@37950
   206
*)
neuper@37950
   207
(*val norm_rational_p : 19.10.02 missing FIXXXXXXXXXXXXME
neuper@37950
   208
      rls               (*.normalizes an rational term [2] without
neuper@37950
   209
                           roots into a simple and canceled fraction
neuper@37950
   210
                           with normalform [2].*)
neuper@37950
   211
*)
neuper@37950
   212
  val norm_rational_ : theory -> term -> (term * term list) option
neuper@37950
   213
  val polynomial2expanded : term -> term option
neuper@37950
   214
  val rational_erls : 
neuper@37950
   215
      rls             (*.evaluates an arbitrary rational term with numerals.*)
neuper@37950
   216
neuper@37950
   217
(*WN0210???SK: fehlen Funktionen, die exportiert werden sollen ? *)
neuper@42451
   218
end (* sig *)
neuper@37950
   219
neuper@37950
   220
(*.**************************************************************************
neuper@37950
   221
survey on the functions
neuper@37950
   222
~~~~~~~~~~~~~~~~~~~~~~~
neuper@37950
   223
 [2] 'polynomial'   :rls               | [3]'expanded_term':rls
neuper@37950
   224
--------------------:------------------+-------------------:-----------------
neuper@37950
   225
 factout_p_         :                  | factout_          :
neuper@37950
   226
 cancel_p_          :                  | cancel_           :
neuper@37950
   227
                    :cancel_p          |                   :cancel
neuper@37950
   228
--------------------:------------------+-------------------:-----------------
neuper@37950
   229
 common_nominator_p_:                  | common_nominator_ :
neuper@37950
   230
                    :common_nominator_p|                   :common_nominator
neuper@37950
   231
 add_fraction_p_    :                  | add_fraction_     :
neuper@37950
   232
--------------------:------------------+-------------------:-----------------
neuper@37950
   233
???SK                 :norm_rational_p   |                   :norm_rational
neuper@37950
   234
neuper@37950
   235
This survey shows only the principal functions for reuse, and the identifiers 
neuper@37950
   236
of the rls exported. The list below shows some more useful functions.
neuper@37950
   237
neuper@37950
   238
neuper@37950
   239
conversion from Isabelle-term to internal representation
neuper@37950
   240
~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
neuper@37950
   241
neuper@37950
   242
... BITTE FORTSETZEN ...
neuper@37950
   243
neuper@37950
   244
polynomial2expanded = ...
neuper@37950
   245
expanded2polynomial = ...
neuper@37950
   246
neuper@37950
   247
remark: polynomial2expanded o expanded2polynomial = I, 
neuper@37950
   248
        where 'o' is function chaining, and 'I' is identity WN0210???SK
neuper@37950
   249
neuper@37950
   250
functions for greatest common divisor and canceling
neuper@37950
   251
~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
neuper@38036
   252
################################################################################
neuper@38036
   253
##   search Isabelle2009-2/src/HOL/Multivariate_Analysis
neuper@38036
   254
##   Amine Chaieb, Robert Himmelmann, John Harrison.
neuper@38036
   255
################################################################################
neuper@37950
   256
mv_gcd
neuper@37950
   257
factout_
neuper@37950
   258
factout_p_
neuper@37950
   259
cancel_
neuper@37950
   260
cancel_p_
neuper@37950
   261
neuper@37950
   262
functions for least common multiple and addition of fractions
neuper@37950
   263
~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
neuper@37950
   264
mv_lcm
neuper@37950
   265
common_nominator_
neuper@37950
   266
common_nominator_p_
neuper@37950
   267
add_fraction_       (*.add 2 or more fractions.*)
neuper@37950
   268
add_fraction_p_     (*.add 2 or more fractions.*)
neuper@37950
   269
neuper@37950
   270
functions for normalform of rationals
neuper@37950
   271
~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
neuper@37950
   272
WN0210???SK interne Funktionen f"ur norm_rational: 
neuper@37950
   273
          schaffen diese SML-Funktionen wirklich ganz allgemeine Terme ?
neuper@37950
   274
neuper@37950
   275
norm_rational_
neuper@37950
   276
norm_expanded_rat_
neuper@37950
   277
neuper@37950
   278
**************************************************************************.*)
neuper@37950
   279
neuper@37950
   280
neuper@37950
   281
(*##*)
neuper@37950
   282
structure RationalI : RATIONALI = 
neuper@37950
   283
struct 
neuper@37950
   284
(*##*)
neuper@37950
   285
neuper@37950
   286
infix mem ins union; (*WN100819 updating to Isabelle2009-2*)
neuper@37950
   287
fun x mem [] = false
neuper@37950
   288
  | x mem (y :: ys) = x = y orelse x mem ys;
neuper@37950
   289
fun (x ins xs) = if x mem xs then xs else x :: xs;
neuper@37950
   290
fun xs union [] = xs
neuper@37950
   291
  | [] union ys = ys
neuper@37950
   292
  | (x :: xs) union ys = xs union (x ins ys);
neuper@37950
   293
neuper@37950
   294
(*. gcd of integers .*)
neuper@37950
   295
(* die gcd Funktion von Isabelle funktioniert nicht richtig !!! *)
neuper@37950
   296
fun gcd_int a b = if b=0 then a
neuper@37950
   297
		  else gcd_int b (a mod b);
neuper@37950
   298
neuper@37950
   299
(*. univariate polynomials (uv) .*)
neuper@37950
   300
(*. univariate polynomials are represented as a list 
neuper@37950
   301
    of the coefficent in reverse maximum degree order .*)
neuper@37950
   302
(*. 5 * x^5 + 4 * x^3 + 2 * x^2 + x + 19 => [19,1,2,4,0,5] .*)
neuper@37950
   303
type uv_poly = int list;
neuper@37950
   304
neuper@37950
   305
(*. adds two uv polynomials .*)
neuper@37950
   306
fun uv_mod_add_poly ([]:uv_poly,p2:uv_poly) = p2:uv_poly 
neuper@37950
   307
  | uv_mod_add_poly (p1,[]) = p1
neuper@37950
   308
  | uv_mod_add_poly (x::p1,y::p2) = (x+y)::(uv_mod_add_poly(p1,p2)); 
neuper@37950
   309
neuper@37950
   310
(*. multiplies a uv polynomial with a skalar s .*)
neuper@37950
   311
fun uv_mod_smul_poly ([]:uv_poly,s:int) = []:uv_poly 
neuper@37950
   312
  | uv_mod_smul_poly (x::p,s) = (x*s)::(uv_mod_smul_poly(p,s)); 
neuper@37950
   313
neuper@37950
   314
(*. calculates the remainder of a polynomial divided by a skalar s .*)
neuper@37950
   315
fun uv_mod_rem_poly ([]:uv_poly,s) = []:uv_poly 
neuper@37950
   316
  | uv_mod_rem_poly (x::p,s) = (x mod s)::(uv_mod_smul_poly(p,s)); 
neuper@37950
   317
neuper@37950
   318
(*. calculates the degree of a uv polynomial .*)
neuper@37950
   319
fun uv_mod_deg ([]:uv_poly) = 0  
neuper@37950
   320
  | uv_mod_deg p = length(p)-1;
neuper@37950
   321
neuper@37950
   322
(*. calculates the remainder of x/p and represents it as 
neuper@37950
   323
    value between -p/2 and p/2 .*)
neuper@37950
   324
fun uv_mod_mod2(x,p)=
neuper@37950
   325
    let
neuper@37950
   326
	val y=(x mod p);
neuper@37950
   327
    in
neuper@37950
   328
	if (y)>(p div 2) then (y)-p else 
neuper@37950
   329
	    (
neuper@37950
   330
	     if (y)<(~p div 2) then p+(y) else (y)
neuper@37950
   331
	     )
neuper@37950
   332
    end;
neuper@37950
   333
neuper@37950
   334
(*.calculates the remainder for each element of a integer list divided by p.*)  
neuper@37950
   335
fun uv_mod_list_modp [] p = [] 
neuper@37950
   336
  | uv_mod_list_modp (x::xs) p = (uv_mod_mod2(x,p))::(uv_mod_list_modp xs p);
neuper@37950
   337
neuper@37950
   338
(*. appends an integer at the end of a integer list .*)
neuper@37950
   339
fun uv_mod_null (p1:int list,0) = p1 
neuper@37950
   340
  | uv_mod_null (p1:int list,n1:int) = uv_mod_null(p1,n1-1) @ [0];
neuper@37950
   341
neuper@37950
   342
(*. uv polynomial division, result is (quotient, remainder) .*)
neuper@37950
   343
(*. only for uv_mod_divides .*)
neuper@37950
   344
(* FIXME: Division von x^9+x^5+1 durch x-1000 funktioniert nicht,
neuper@37950
   345
   integer zu klein  *)
neuper@37950
   346
fun uv_mod_pdiv (p1:uv_poly) ([]:uv_poly) = 
neuper@38031
   347
    error ("RATIONALS_UV_MOD_PDIV_EXCEPTION: division by zero")
neuper@37950
   348
  | uv_mod_pdiv p1 [x] = 
neuper@37950
   349
    let
neuper@38006
   350
	val xs= Unsynchronized.ref  [];
neuper@37950
   351
    in
neuper@37950
   352
	if x<>0 then 
neuper@37950
   353
	    (
neuper@37950
   354
	     xs:=(uv_mod_rem_poly(p1,x));
neuper@37950
   355
	     while length(!xs)>0 andalso hd(!xs)=0 do xs:=tl(!xs)
neuper@37950
   356
	     )
neuper@38031
   357
	else error ("RATIONALS_UV_MOD_PDIV_EXCEPTION: division by zero");
neuper@37950
   358
	([]:uv_poly,!xs:uv_poly)
neuper@37950
   359
    end
neuper@37950
   360
  | uv_mod_pdiv p1 p2 =  
neuper@37950
   361
    let
neuper@37950
   362
	val n= uv_mod_deg(p2);
neuper@38006
   363
	val m= Unsynchronized.ref (uv_mod_deg(p1));
neuper@38006
   364
	val p1'= Unsynchronized.ref  (rev(p1));
neuper@37950
   365
	val p2'=(rev(p2));
neuper@37950
   366
	val lc2=hd(p2');
neuper@38006
   367
	val q= Unsynchronized.ref  [];
neuper@38006
   368
	val c= Unsynchronized.ref  0;
neuper@38006
   369
	val output= Unsynchronized.ref  ([],[]);
neuper@37950
   370
    in
neuper@37950
   371
	(
neuper@37950
   372
	 if (!m)=0 orelse p2=[0] 
neuper@38031
   373
         then error ("RATIONALS_UV_MOD_PDIV_EXCEPTION: Division by zero") 
neuper@37950
   374
	 else
neuper@37950
   375
	     (
neuper@37950
   376
	      if (!m)<n then 
neuper@37950
   377
		  (
neuper@37950
   378
		   output:=([0],p1) 
neuper@37950
   379
		   ) 
neuper@37950
   380
	      else
neuper@37950
   381
		  (
neuper@37950
   382
		   while (!m)>=n do
neuper@37950
   383
		       (
neuper@37950
   384
			c:=hd(!p1') div hd(p2');
neuper@37950
   385
			if !c<>0 then
neuper@37950
   386
			    (
neuper@37950
   387
			     p1':=uv_mod_add_poly(!p1',uv_mod_null(uv_mod_smul_poly(p2',~(!c)),!m-n));
neuper@37950
   388
			     while length(!p1')>0 andalso hd(!p1')=0  do p1':= tl(!p1');
neuper@37950
   389
			     m:=uv_mod_deg(!p1')
neuper@37950
   390
			     )
neuper@37950
   391
			else m:=0
neuper@37950
   392
			);
neuper@37950
   393
    		   output:=(rev(!q),rev(!p1'))
neuper@37950
   394
		   )
neuper@37950
   395
	      );
neuper@37950
   396
	     !output
neuper@37950
   397
	 )
neuper@37950
   398
    end;
neuper@37950
   399
neuper@37950
   400
(*. divides p1 by p2 in Zp .*)
neuper@37950
   401
fun uv_mod_pdivp (p1:uv_poly) (p2:uv_poly) p =  
neuper@37950
   402
    let
neuper@37950
   403
	val n=uv_mod_deg(p2);
neuper@38006
   404
	val m= Unsynchronized.ref  (uv_mod_deg(uv_mod_list_modp p1 p));
neuper@38006
   405
	val p1'= Unsynchronized.ref  (rev(p1));
neuper@37950
   406
	val p2'=(rev(uv_mod_list_modp p2 p));
neuper@37950
   407
	val lc2=hd(p2');
neuper@38006
   408
	val q= Unsynchronized.ref  [];
neuper@38006
   409
	val c= Unsynchronized.ref  0;
neuper@38006
   410
	val output= Unsynchronized.ref  ([],[]);
neuper@37950
   411
    in
neuper@37950
   412
	(
neuper@38031
   413
	 if (!m)=0 orelse p2=[0] then error ("RATIONALS_UV_MOD_PDIVP_EXCEPTION: Division by zero") 
neuper@37950
   414
	 else
neuper@37950
   415
	     (
neuper@37950
   416
	      if (!m)<n then 
neuper@37950
   417
		  (
neuper@37950
   418
		   output:=([0],p1) 
neuper@37950
   419
		   ) 
neuper@37950
   420
	      else
neuper@37950
   421
		  (
neuper@37950
   422
		   while (!m)>=n do
neuper@37950
   423
		       (
neuper@37950
   424
			c:=uv_mod_mod2(hd(!p1')*(power lc2 1), p);
neuper@37950
   425
			q:=(!c)::(!q);
neuper@37950
   426
			p1':=uv_mod_list_modp(tl(uv_mod_add_poly(uv_mod_smul_poly(!p1',lc2),
neuper@37950
   427
								  uv_mod_smul_poly(uv_mod_smul_poly(p2',hd(!p1')),~1)))) p;
neuper@37950
   428
			m:=(!m)-1
neuper@37950
   429
			);
neuper@37950
   430
		   
neuper@37950
   431
		   while !p1'<>[] andalso hd(!p1')=0 do
neuper@37950
   432
		       (
neuper@37950
   433
			p1':=tl(!p1')
neuper@37950
   434
			);
neuper@37950
   435
neuper@37950
   436
    		   output:=(rev(uv_mod_list_modp (!q) (p)),rev(!p1'))
neuper@37950
   437
		   )
neuper@37950
   438
	      );
neuper@37950
   439
	     !output:uv_poly * uv_poly
neuper@37950
   440
	 )
neuper@37950
   441
    end;
neuper@37950
   442
neuper@37950
   443
(*. calculates the remainder of p1/p2 .*)
neuper@38031
   444
fun uv_mod_prest (p1:uv_poly) ([]:uv_poly) = error("UV_MOD_PREST_EXCEPTION: Division by zero") 
neuper@37950
   445
  | uv_mod_prest [] p2 = []:uv_poly
neuper@37950
   446
  | uv_mod_prest p1 p2 = (#2(uv_mod_pdiv p1 p2));
neuper@37950
   447
neuper@37950
   448
(*. calculates the remainder of p1/p2 in Zp .*)
neuper@38031
   449
fun uv_mod_prestp (p1:uv_poly) ([]:uv_poly) p= error("UV_MOD_PRESTP_EXCEPTION: Division by zero") 
neuper@37950
   450
  | uv_mod_prestp [] p2 p= []:uv_poly 
neuper@37950
   451
  | uv_mod_prestp p1 p2 p = #2(uv_mod_pdivp p1 p2 p); 
neuper@37950
   452
neuper@37950
   453
(*. calculates the content of a uv polynomial .*)
neuper@37950
   454
fun uv_mod_cont ([]:uv_poly) = 0  
neuper@37950
   455
  | uv_mod_cont (x::p)= gcd_int x (uv_mod_cont(p));
neuper@37950
   456
neuper@37950
   457
(*. divides each coefficient of a uv polynomial by y .*)
neuper@38031
   458
fun uv_mod_div_list (p:uv_poly,0) = error("UV_MOD_DIV_LIST_EXCEPTION: Division by zero") 
neuper@37950
   459
  | uv_mod_div_list ([],y)   = []:uv_poly
neuper@37950
   460
  | uv_mod_div_list (x::p,y) = (x div y)::uv_mod_div_list(p,y); 
neuper@37950
   461
neuper@37950
   462
(*. calculates the primitiv part of a uv polynomial .*)
neuper@37950
   463
fun uv_mod_pp ([]:uv_poly) = []:uv_poly
neuper@37950
   464
  | uv_mod_pp p =  
neuper@37950
   465
    let
neuper@38006
   466
	val c= Unsynchronized.ref  0;
neuper@37950
   467
    in
neuper@37950
   468
	(
neuper@37950
   469
	 c:=uv_mod_cont(p);
neuper@37950
   470
	 
neuper@38031
   471
	 if !c=0 then error ("RATIONALS_UV_MOD_PP_EXCEPTION: content is 0")
neuper@37950
   472
	 else uv_mod_div_list(p,!c)
neuper@37950
   473
	)
neuper@37950
   474
    end;
neuper@37950
   475
neuper@37950
   476
(*. gets the leading coefficient of a uv polynomial .*)
neuper@37950
   477
fun uv_mod_lc ([]:uv_poly) = 0 
neuper@37950
   478
  | uv_mod_lc p  = hd(rev(p)); 
neuper@37950
   479
neuper@37950
   480
(*. calculates the euklidean polynomial remainder sequence in Zp .*)
neuper@37950
   481
fun uv_mod_prs_euklid_p(p1:uv_poly,p2:uv_poly,p)= 
neuper@37950
   482
    let
neuper@38006
   483
	val f = Unsynchronized.ref  [];
neuper@38006
   484
	val f'= Unsynchronized.ref  p2;
neuper@38006
   485
	val fi= Unsynchronized.ref  [];
neuper@37950
   486
    in
neuper@37950
   487
	( 
neuper@37950
   488
	 f:=p2::p1::[]; 
neuper@37950
   489
 	 while uv_mod_deg(!f')>0 do
neuper@37950
   490
	     (
neuper@37950
   491
	      f':=uv_mod_prestp (hd(tl(!f))) (hd(!f)) p;
neuper@37950
   492
	      if (!f')<>[] then 
neuper@37950
   493
		  (
neuper@37950
   494
		   fi:=(!f');
neuper@37950
   495
		   f:=(!fi)::(!f)
neuper@37950
   496
		   )
neuper@37950
   497
	      else ()
neuper@37950
   498
	      );
neuper@37950
   499
	      (!f)
neuper@37950
   500
	 
neuper@37950
   501
	 )
neuper@37950
   502
    end;
neuper@37950
   503
neuper@37950
   504
(*. calculates the gcd of p1 and p2 in Zp .*)
neuper@37950
   505
fun uv_mod_gcd_modp ([]:uv_poly) (p2:uv_poly) p = p2:uv_poly 
neuper@37950
   506
  | uv_mod_gcd_modp p1 [] p= p1
neuper@37950
   507
  | uv_mod_gcd_modp p1 p2 p=
neuper@37950
   508
    let
neuper@38006
   509
	val p1'= Unsynchronized.ref [];
neuper@38006
   510
	val p2'= Unsynchronized.ref [];
neuper@38006
   511
	val pc= Unsynchronized.ref [];
neuper@38006
   512
	val g= Unsynchronized.ref  [];
neuper@38006
   513
	val d= Unsynchronized.ref  0;
neuper@38006
   514
	val prs= Unsynchronized.ref  [];
neuper@37950
   515
    in
neuper@37950
   516
	(
neuper@37950
   517
	 if uv_mod_deg(p1)>=uv_mod_deg(p2) then
neuper@37950
   518
	     (
neuper@37950
   519
	      p1':=uv_mod_list_modp (uv_mod_pp(p1)) p;
neuper@37950
   520
	      p2':=uv_mod_list_modp (uv_mod_pp(p2)) p
neuper@37950
   521
	      )
neuper@37950
   522
	 else 
neuper@37950
   523
	     (
neuper@37950
   524
	      p1':=uv_mod_list_modp (uv_mod_pp(p2)) p;
neuper@37950
   525
	      p2':=uv_mod_list_modp (uv_mod_pp(p1)) p
neuper@37950
   526
	      );
neuper@37950
   527
	 d:=uv_mod_mod2((gcd_int (uv_mod_cont(p1))) (uv_mod_cont(p2)), p) ;
neuper@37950
   528
	 if !d>(p div 2) then d:=(!d)-p else ();
neuper@37950
   529
	 
neuper@37950
   530
	 prs:=uv_mod_prs_euklid_p(!p1',!p2',p);
neuper@37950
   531
neuper@37950
   532
	 if hd(!prs)=[] then pc:=hd(tl(!prs))
neuper@37950
   533
	 else pc:=hd(!prs);
neuper@37950
   534
neuper@37950
   535
	 g:=uv_mod_smul_poly(uv_mod_pp(!pc),!d);
neuper@37950
   536
	 !g
neuper@37950
   537
	 )
neuper@37950
   538
    end;
neuper@37950
   539
neuper@37950
   540
(*. calculates the minimum of two real values x and y .*)
neuper@37978
   541
fun uv_mod_r_min(x,y):Real.real = if x>y then y else x;
neuper@37950
   542
neuper@37950
   543
(*. calculates the minimum of two integer values x and y .*)
neuper@37950
   544
fun uv_mod_min(x,y) = if x>y then y else x;
neuper@37950
   545
neuper@37950
   546
(*. adds the squared values of a integer list .*)
neuper@37950
   547
fun uv_mod_add_qu [] = 0.0 
neuper@37978
   548
  | uv_mod_add_qu (x::p) =  Real.fromInt(x)*Real.fromInt(x) + uv_mod_add_qu p;
neuper@37950
   549
neuper@37950
   550
(*. calculates the euklidean norm .*)
neuper@37950
   551
fun uv_mod_norm ([]:uv_poly) = 0.0
neuper@37950
   552
  | uv_mod_norm p = Math.sqrt(uv_mod_add_qu(p));
neuper@37950
   553
neuper@37950
   554
(*. multipies two values a and b .*)
neuper@37950
   555
fun uv_mod_multi a b = a * b;
neuper@37950
   556
neuper@37950
   557
(*. decides if x is a prim, the list contains all primes which are lower then x .*)
neuper@37950
   558
fun uv_mod_prim(x,[])= false 
neuper@37950
   559
  | uv_mod_prim(x,[y])=if ((x mod y) <> 0) then true
neuper@37950
   560
		else false
neuper@37950
   561
  | uv_mod_prim(x,y::ys) = if uv_mod_prim(x,[y])
neuper@37950
   562
			then 
neuper@37950
   563
			    if uv_mod_prim(x,ys) then true 
neuper@37950
   564
			    else false
neuper@37950
   565
		    else false;
neuper@37950
   566
neuper@37950
   567
(*. gets the first prime, which is greater than p and does not divide g .*)
neuper@37950
   568
fun uv_mod_nextprime(g,p)= 
neuper@37950
   569
    let
neuper@38006
   570
	val list= Unsynchronized.ref  [2];
neuper@38006
   571
	val exit= Unsynchronized.ref  0;
neuper@38006
   572
	val i = Unsynchronized.ref 2
neuper@37950
   573
    in
neuper@37950
   574
	while (!i<p) do (* calculates the primes lower then p *)
neuper@37950
   575
	    (
neuper@37950
   576
	     if uv_mod_prim(!i,!list) then
neuper@37950
   577
		 (
neuper@37950
   578
		  if (p mod !i <> 0)
neuper@37950
   579
		      then
neuper@37950
   580
			  (
neuper@37950
   581
			   list:= (!i)::(!list);
neuper@37950
   582
			   i:= (!i)+1
neuper@37950
   583
			   )
neuper@37950
   584
		  else i:=(!i)+1
neuper@37950
   585
		  )
neuper@37950
   586
	     else i:= (!i)+1
neuper@37950
   587
		 );
neuper@37950
   588
	    i:=(p+1);
neuper@37950
   589
	    while (!exit=0) do   (* calculate next prime which does not divide g *)
neuper@37950
   590
	    (
neuper@37950
   591
	     if uv_mod_prim(!i,!list) then
neuper@37950
   592
		 (
neuper@37950
   593
		  if (g mod !i <> 0)
neuper@37950
   594
		      then
neuper@37950
   595
			  (
neuper@37950
   596
			   list:= (!i)::(!list);
neuper@37950
   597
			   exit:= (!i)
neuper@37950
   598
			   )
neuper@37950
   599
		  else i:=(!i)+1
neuper@37950
   600
		      )
neuper@37950
   601
	     else i:= (!i)+1
neuper@37950
   602
		 ); 
neuper@37950
   603
	    !exit
neuper@37950
   604
    end;
neuper@37950
   605
neuper@37950
   606
(*. decides if p1 is a factor of p2 in Zp .*)
neuper@38031
   607
fun uv_mod_dividesp ([]:uv_poly) (p2:uv_poly) p= error("UV_MOD_DIVIDESP: Division by zero") 
neuper@37950
   608
  | uv_mod_dividesp p1 p2 p= if uv_mod_prestp p2 p1 p = [] then true else false;
neuper@37950
   609
neuper@37950
   610
(*. decides if p1 is a factor of p2 .*)
neuper@38031
   611
fun uv_mod_divides ([]:uv_poly) (p2:uv_poly) = error("UV_MOD_DIVIDES: Division by zero")
neuper@37950
   612
  | uv_mod_divides p1 p2 = if uv_mod_prest p2 p1  = [] then true else false;
neuper@37950
   613
neuper@37950
   614
(*. chinese remainder algorithm .*)
neuper@37950
   615
fun uv_mod_cra2(r1,r2,m1,m2)=     
neuper@37950
   616
    let 
neuper@38006
   617
	val c= Unsynchronized.ref  0;
neuper@38006
   618
	val r1'= Unsynchronized.ref  0;
neuper@38006
   619
	val d= Unsynchronized.ref  0;
neuper@38006
   620
	val a= Unsynchronized.ref  0;
neuper@37950
   621
    in
neuper@37950
   622
	(
neuper@37950
   623
	 while (uv_mod_mod2((!c)*m1,m2))<>1 do 
neuper@37950
   624
	     (
neuper@37950
   625
	      c:=(!c)+1
neuper@37950
   626
	      );
neuper@37950
   627
	 r1':= uv_mod_mod2(r1,m1);
neuper@37950
   628
	 d:=uv_mod_mod2(((r2-(!r1'))*(!c)),m2);
neuper@37950
   629
	 !r1'+(!d)*m1    
neuper@37950
   630
	 )
neuper@37950
   631
    end;
neuper@37950
   632
neuper@37950
   633
(*. applies the chinese remainder algorithmen to the coefficients of x1 and x2 .*)
neuper@37950
   634
fun uv_mod_cra_2 ([],[],m1,m2) = [] 
neuper@38031
   635
  | uv_mod_cra_2 ([],x2,m1,m2) = error("UV_MOD_CRA_2_EXCEPTION: invalid call x1")
neuper@38031
   636
  | uv_mod_cra_2 (x1,[],m1,m2) = error("UV_MOD_CRA_2_EXCEPTION: invalid call x2")
neuper@37950
   637
  | uv_mod_cra_2 (x1::x1s,x2::x2s,m1,m2) = (uv_mod_cra2(x1,x2,m1,m2))::(uv_mod_cra_2(x1s,x2s,m1,m2));
neuper@37950
   638
neuper@37950
   639
(*. calculates the gcd of two uv polynomials p1' and p2' with the modular algorithm .*)
neuper@37950
   640
fun uv_mod_gcd (p1':uv_poly) (p2':uv_poly) =
neuper@37950
   641
    let 
neuper@38006
   642
	val p1= Unsynchronized.ref  (uv_mod_pp(p1'));
neuper@38006
   643
	val p2= Unsynchronized.ref  (uv_mod_pp(p2'));
neuper@37950
   644
	val c=gcd_int (uv_mod_cont(p1')) (uv_mod_cont(p2'));
neuper@38006
   645
	val temp= Unsynchronized.ref  [];
neuper@38006
   646
	val cp= Unsynchronized.ref  [];
neuper@38006
   647
	val qp= Unsynchronized.ref  [];
neuper@38006
   648
	val q= Unsynchronized.ref [];
neuper@38006
   649
	val pn= Unsynchronized.ref  0;
neuper@38006
   650
	val d= Unsynchronized.ref  0;
neuper@38006
   651
	val g1= Unsynchronized.ref  0;
neuper@38006
   652
	val p= Unsynchronized.ref  0;    
neuper@38006
   653
	val m= Unsynchronized.ref  0;
neuper@38006
   654
	val exit= Unsynchronized.ref  0;
neuper@38006
   655
	val i= Unsynchronized.ref  1;
neuper@37950
   656
    in
neuper@37950
   657
	if length(!p1)>length(!p2) then ()
neuper@37950
   658
	else 
neuper@37950
   659
	    (
neuper@37950
   660
	     temp:= !p1;
neuper@37950
   661
	     p1:= !p2;
neuper@37950
   662
	     p2:= !temp
neuper@37950
   663
	     );
neuper@37950
   664
neuper@37950
   665
	 
neuper@37950
   666
	d:=gcd_int (uv_mod_lc(!p1)) (uv_mod_lc(!p2));
neuper@37950
   667
	g1:=uv_mod_lc(!p1)*uv_mod_lc(!p2);
neuper@37950
   668
	p:=4;
neuper@37950
   669
	
neuper@37978
   670
	m:=Real.ceil(2.0 * Real.fromInt(!d) *
neuper@37978
   671
	  Real.fromInt(power 2 (uv_mod_min(uv_mod_deg(!p2),uv_mod_deg(!p1)))) *
neuper@37978
   672
	  Real.fromInt(!d) * 
neuper@37978
   673
	  uv_mod_r_min(uv_mod_norm(!p1) / Real.fromInt(abs(uv_mod_lc(!p1))),
neuper@37978
   674
	  uv_mod_norm(!p2) / Real.fromInt(abs(uv_mod_lc(!p2))))); 
neuper@37950
   675
neuper@37950
   676
	while (!exit=0) do  
neuper@37950
   677
	    (
neuper@37950
   678
	     p:=uv_mod_nextprime(!d,!p);
neuper@37950
   679
	     cp:=(uv_mod_gcd_modp (uv_mod_list_modp(!p1) (!p)) (uv_mod_list_modp(!p2) (!p)) (!p)) ;
neuper@37950
   680
	     if abs(uv_mod_lc(!cp))<>1 then  (* leading coefficient = 1 ? *)
neuper@37950
   681
		 (
neuper@37950
   682
		  i:=1;
neuper@37950
   683
		  while (!i)<(!p) andalso (abs(uv_mod_mod2((uv_mod_lc(!cp)*(!i)),(!p)))<>1) do
neuper@37950
   684
		      (
neuper@37950
   685
		       i:=(!i)+1
neuper@37950
   686
		       );
neuper@37950
   687
		      cp:=uv_mod_list_modp (map (uv_mod_multi (!i)) (!cp)) (!p) 
neuper@37950
   688
		  )
neuper@37950
   689
	     else ();
neuper@37950
   690
neuper@37950
   691
	     qp:= ((map (uv_mod_multi (uv_mod_mod2(!d,!p)))) (!cp));
neuper@37950
   692
neuper@37950
   693
	     if uv_mod_deg(!qp)=0 then (q:=[1]; exit:=1) else ();
neuper@37950
   694
neuper@37950
   695
	     pn:=(!p);
neuper@37950
   696
	     q:=(!qp);
neuper@37950
   697
neuper@37950
   698
	     while !pn<= !m andalso !m>(!p) andalso !exit=0 do
neuper@37950
   699
		 (
neuper@37950
   700
		  p:=uv_mod_nextprime(!d,!p);
neuper@37950
   701
 		  cp:=(uv_mod_gcd_modp (uv_mod_list_modp(!p1) (!p)) (uv_mod_list_modp(!p2) (!p)) (!p)); 
neuper@37950
   702
 		  if uv_mod_lc(!cp)<>1 then  (* leading coefficient = 1 ? *)
neuper@37950
   703
 		      (
neuper@37950
   704
 		       i:=1;
neuper@37950
   705
 		       while (!i)<(!p) andalso ((uv_mod_mod2((uv_mod_lc(!q)*(!i)),(!p)))<>1) do
neuper@37950
   706
 			   (
neuper@37950
   707
 			    i:=(!i)+1
neuper@37950
   708
		           );
neuper@37950
   709
		       cp:=uv_mod_list_modp (map (uv_mod_multi (!i)) (!cp)) (!p)
neuper@37950
   710
 		      )
neuper@37950
   711
 		  else ();    
neuper@37950
   712
 		 
neuper@37950
   713
		  qp:=uv_mod_list_modp ((map (uv_mod_multi (uv_mod_mod2(!d,!p)))) (!cp)  ) (!p);
neuper@37950
   714
 		  if uv_mod_deg(!qp)>uv_mod_deg(!q) then
neuper@37950
   715
 		      (
neuper@37950
   716
 		       (*print("degree to high!!!\n")*)
neuper@37950
   717
 		       )
neuper@37950
   718
 		  else
neuper@37950
   719
 		      (
neuper@37950
   720
 		       if uv_mod_deg(!qp)=uv_mod_deg(!q) then
neuper@37950
   721
 			   (
neuper@37950
   722
 			    q:=uv_mod_cra_2(!q,!qp,!pn,!p);
neuper@37950
   723
			    pn:=(!pn) * !p;
neuper@37950
   724
			    q:=uv_mod_pp(uv_mod_list_modp (!q) (!pn)); (* found already gcd ? *)
neuper@37950
   725
			    if (uv_mod_divides (!q) (p1')) andalso (uv_mod_divides (!q) (p2')) then (exit:=1) else ()
neuper@37950
   726
		 	    )
neuper@37950
   727
		       else
neuper@37950
   728
			   (
neuper@37950
   729
			    if  uv_mod_deg(!qp)<uv_mod_deg(!q) then
neuper@37950
   730
				(
neuper@37950
   731
				 pn:= !p;
neuper@37950
   732
				 q:= !qp
neuper@37950
   733
				 )
neuper@37950
   734
			    else ()
neuper@37950
   735
			    )
neuper@37950
   736
		       )
neuper@37950
   737
		  );
neuper@37950
   738
 	     q:=uv_mod_pp(uv_mod_list_modp (!q) (!pn));
neuper@37950
   739
	     if (uv_mod_divides (!q) (p1')) andalso (uv_mod_divides (!q) (p2')) then exit:=1 else ()
neuper@37950
   740
	     );
neuper@37950
   741
	    uv_mod_smul_poly(!q,c):uv_poly
neuper@37950
   742
    end;
neuper@37950
   743
neuper@37950
   744
(*. multivariate polynomials .*)
neuper@37950
   745
(*. multivariate polynomials are represented as a list of the pairs, 
neuper@37950
   746
 first is the coefficent and the second is a list of the exponents .*)
neuper@37950
   747
(*. 5 * x^5 * y^3 + 4 * x^3 * z^2 + 2 * x^2 * y * z^3 - z - 19 
neuper@37950
   748
 => [(5,[5,3,0]),(4,[3,0,2]),(2,[2,1,3]),(~1,[0,0,1]),(~19,[0,0,0])] .*)
neuper@37950
   749
neuper@37950
   750
(*. global variables .*)
neuper@37950
   751
(*. order indicators .*)
neuper@37950
   752
val LEX_=0; (* lexicographical term order *)
neuper@37950
   753
val GGO_=1; (* greatest degree order *)
neuper@37950
   754
neuper@37950
   755
(*. datatypes for internal representation.*)
neuper@37950
   756
type mv_monom = (int *      (*.coefficient or the monom.*)
neuper@37950
   757
		 int list); (*.list of exponents)      .*)
neuper@37950
   758
fun mv_monom2str (i, is) = "("^ int2str i^"," ^ ints2str' is ^ ")";
neuper@37950
   759
neuper@37950
   760
type mv_poly = mv_monom list; 
neuper@37950
   761
fun mv_poly2str p = (strs2str' o (map mv_monom2str)) p;
neuper@37950
   762
neuper@37950
   763
(*. help function for monom_greater and geq .*)
neuper@37950
   764
fun mv_mg_hlp([]) = EQUAL 
neuper@37950
   765
  | mv_mg_hlp(x::list)=if x<0 then LESS
neuper@37950
   766
		    else if x>0 then GREATER
neuper@37950
   767
			 else mv_mg_hlp(list);
neuper@37950
   768
neuper@37950
   769
(*. adds a list of values .*)
neuper@37950
   770
fun mv_addlist([]) = 0
neuper@37950
   771
  | mv_addlist(p1) = hd(p1)+mv_addlist(tl(p1));
neuper@37950
   772
			   
neuper@37950
   773
(*. tests if the monomial M1 is greater as the monomial M2 and returns a boolean value .*)
neuper@37950
   774
(*. 2 orders are implemented LEX_/GGO_ (lexigraphical/greatest degree order) .*)
neuper@37950
   775
fun mv_monom_greater((M1x,M1l):mv_monom,(M2x,M2l):mv_monom,order)=
neuper@37950
   776
    if order=LEX_ then
neuper@37950
   777
	( 
neuper@38031
   778
	 if length(M1l)<>length(M2l) then error ("RATIONALS_MV_MONOM_GREATER_EXCEPTION: Order error")
neuper@37950
   779
	 else if (mv_mg_hlp((map op- (M1l~~M2l)))<>GREATER) then false else true
neuper@37950
   780
	     )
neuper@37950
   781
    else
neuper@37950
   782
	if order=GGO_ then
neuper@37950
   783
	    ( 
neuper@38031
   784
	     if length(M1l)<>length(M2l) then error ("RATIONALS_MV_MONOM_GREATER_EXCEPTION: Order error")
neuper@37950
   785
	     else 
neuper@37950
   786
		 if mv_addlist(M1l)=mv_addlist(M2l)  then if (mv_mg_hlp((map op- (M1l~~M2l)))<>GREATER) then false else true
neuper@37950
   787
		 else if mv_addlist(M1l)>mv_addlist(M2l) then true else false
neuper@37950
   788
	     )
neuper@38031
   789
	else error ("RATIONALS_MV_MONOM_GREATER_EXCEPTION: Wrong Order");
neuper@37950
   790
		   
neuper@37950
   791
(*. tests if the monomial X is greater as the monomial Y and returns a order value (GREATER,EQUAL,LESS) .*)
neuper@37950
   792
(*. 2 orders are implemented LEX_/GGO_ (lexigraphical/greatest degree order) .*)
neuper@37950
   793
fun mv_geq order ((x1,x):mv_monom,(x2,y):mv_monom) =
neuper@37950
   794
let 
neuper@38006
   795
    val temp= Unsynchronized.ref  EQUAL;
neuper@37950
   796
in
neuper@37950
   797
    if order=LEX_ then
neuper@37950
   798
	(
neuper@37950
   799
	 if length(x)<>length(y) then 
neuper@38031
   800
	     error ("RATIONALS_MV_GEQ_EXCEPTION: Order error")
neuper@37950
   801
	 else 
neuper@37950
   802
	     (
neuper@37950
   803
	      temp:=mv_mg_hlp((map op- (x~~y)));
neuper@37950
   804
	      if !temp=EQUAL then 
neuper@37950
   805
		  ( if x1=x2 then EQUAL 
neuper@37950
   806
		    else if x1>x2 then GREATER
neuper@37950
   807
			 else LESS
neuper@37950
   808
			     )
neuper@37950
   809
	      else (!temp)
neuper@37950
   810
	      )
neuper@37950
   811
	     )
neuper@37950
   812
    else 
neuper@37950
   813
	if order=GGO_ then 
neuper@37950
   814
	    (
neuper@37950
   815
	     if length(x)<>length(y) then 
neuper@38031
   816
	      error ("RATIONALS_MV_GEQ_EXCEPTION: Order error")
neuper@37950
   817
	     else 
neuper@37950
   818
		 if mv_addlist(x)=mv_addlist(y) then 
neuper@37950
   819
		     (mv_mg_hlp((map op- (x~~y))))
neuper@37950
   820
		 else if mv_addlist(x)>mv_addlist(y) then GREATER else LESS
neuper@37950
   821
		     )
neuper@38031
   822
	else error ("RATIONALS_MV_GEQ_EXCEPTION: Wrong Order")
neuper@37950
   823
end;
neuper@37950
   824
neuper@37950
   825
(*. cuts the first variable from a polynomial .*)
neuper@37950
   826
fun mv_cut([]:mv_poly)=[]:mv_poly
neuper@38031
   827
  | mv_cut((x,[])::list) = error ("RATIONALS_MV_CUT_EXCEPTION: Invalid list ")
neuper@37950
   828
  | mv_cut((x,y::ys)::list)=(x,ys)::mv_cut(list);
neuper@37950
   829
	    
neuper@37950
   830
(*. leading power product .*)
neuper@37950
   831
fun mv_lpp([]:mv_poly,order)  = []
neuper@37950
   832
  | mv_lpp([(x,y)],order) = y
neuper@37950
   833
  | mv_lpp(p1,order)  = #2(hd(rev(sort (mv_geq order) p1)));
neuper@37950
   834
    
neuper@37950
   835
(*. leading monomial .*)
neuper@37950
   836
fun mv_lm([]:mv_poly,order)  = (0,[]):mv_monom
neuper@37950
   837
  | mv_lm([x],order) = x 
neuper@37950
   838
  | mv_lm(p1,order)  = hd(rev(sort (mv_geq order) p1));
neuper@37950
   839
    
neuper@37950
   840
(*. leading coefficient in term order .*)
neuper@37950
   841
fun mv_lc2([]:mv_poly,order)  = 0
neuper@37950
   842
  | mv_lc2([(x,y)],order) = x
neuper@37950
   843
  | mv_lc2(p1,order)  = #1(hd(rev(sort (mv_geq order) p1)));
neuper@37950
   844
neuper@37950
   845
neuper@37950
   846
(*. reverse the coefficients in mv polynomial .*)
neuper@37950
   847
fun mv_rev_to([]:mv_poly) = []:mv_poly
neuper@37950
   848
  | mv_rev_to((c,e)::xs) = (c,rev(e))::mv_rev_to(xs);
neuper@37950
   849
neuper@37950
   850
(*. leading coefficient in reverse term order .*)
neuper@37950
   851
fun mv_lc([]:mv_poly,order)  = []:mv_poly 
neuper@37950
   852
  | mv_lc([(x,y)],order) = mv_rev_to(mv_cut(mv_rev_to([(x,y)])))
neuper@37950
   853
  | mv_lc(p1,order)  = 
neuper@37950
   854
    let
neuper@38006
   855
	val p1o= Unsynchronized.ref  (rev(sort (mv_geq order) (mv_rev_to(p1))));
neuper@37950
   856
	val lp=hd(#2(hd(!p1o)));
neuper@38006
   857
	val lc= Unsynchronized.ref  [];
neuper@37950
   858
    in
neuper@37950
   859
	(
neuper@37950
   860
	 while (length(!p1o)>0 andalso hd(#2(hd(!p1o)))=lp) do
neuper@37950
   861
	     (
neuper@37950
   862
	      lc:=hd(mv_cut([hd(!p1o)]))::(!lc);
neuper@37950
   863
	      p1o:=tl(!p1o)
neuper@37950
   864
	      );
neuper@38031
   865
	 if !lc=[] then error ("RATIONALS_MV_LC_EXCEPTION: lc is empty") else ();
neuper@37950
   866
	 mv_rev_to(!lc)
neuper@37950
   867
	 )
neuper@37950
   868
    end;
neuper@37950
   869
neuper@37950
   870
(*. compares two powerproducts .*)
neuper@37950
   871
fun mv_monom_equal((_,xlist):mv_monom,(_,ylist):mv_monom) = (foldr and_) (((map op=) (xlist~~ylist)),true);
neuper@37950
   872
    
neuper@37950
   873
(*. help function for mv_add .*)
neuper@37950
   874
fun mv_madd([]:mv_poly,[]:mv_poly,order) = []:mv_poly
neuper@37950
   875
  | mv_madd([(0,_)],p2,order) = p2
neuper@37950
   876
  | mv_madd(p1,[(0,_)],order) = p1  
neuper@37950
   877
  | mv_madd([],p2,order) = p2
neuper@37950
   878
  | mv_madd(p1,[],order) = p1
neuper@37950
   879
  | mv_madd(p1,p2,order) = 
neuper@37950
   880
    (
neuper@37950
   881
     if mv_monom_greater(hd(p1),hd(p2),order) 
neuper@37950
   882
	 then hd(p1)::mv_madd(tl(p1),p2,order)
neuper@37950
   883
     else if mv_monom_equal(hd(p1),hd(p2)) 
neuper@37950
   884
	      then if mv_lc2(p1,order)+mv_lc2(p2,order)<>0 
neuper@37950
   885
		       then (mv_lc2(p1,order)+mv_lc2(p2,order),mv_lpp(p1,order))::mv_madd(tl(p1),tl(p2),order)
neuper@37950
   886
		   else mv_madd(tl(p1),tl(p2),order)
neuper@37950
   887
	  else hd(p2)::mv_madd(p1,tl(p2),order)
neuper@37950
   888
	      )
neuper@37950
   889
	      
neuper@37950
   890
(*. adds two multivariate polynomials .*)
neuper@37950
   891
fun mv_add([]:mv_poly,p2:mv_poly,order) = p2
neuper@37950
   892
  | mv_add(p1,[],order) = p1
neuper@37950
   893
  | mv_add(p1,p2,order) = mv_madd(rev(sort (mv_geq order) p1),rev(sort (mv_geq order) p2), order);
neuper@37950
   894
neuper@37950
   895
(*. monom multiplication .*)
neuper@37950
   896
fun mv_mmul((x1,y1):mv_monom,(x2,y2):mv_monom)=(x1*x2,(map op+) (y1~~y2)):mv_monom;
neuper@37950
   897
neuper@37950
   898
(*. deletes all monomials with coefficient 0 .*)
neuper@37950
   899
fun mv_shorten([]:mv_poly,order) = []:mv_poly
neuper@37950
   900
  | mv_shorten(x::xs,order)=mv_madd([x],mv_shorten(xs,order),order);
neuper@37950
   901
neuper@37950
   902
(*. zeros a list .*)
neuper@37950
   903
fun mv_null2([])=[]
neuper@37950
   904
  | mv_null2(x::l)=0::mv_null2(l);
neuper@37950
   905
neuper@37950
   906
(*. multiplies two multivariate polynomials .*)
neuper@37950
   907
fun mv_mul([]:mv_poly,[]:mv_poly,_) = []:mv_poly
neuper@37950
   908
  | mv_mul([],y::p2,_) = [(0,mv_null2(#2(y)))]
neuper@37950
   909
  | mv_mul(x::p1,[],_) = [(0,mv_null2(#2(x)))] 
neuper@37950
   910
  | mv_mul(x::p1,y::p2,order) = mv_shorten(rev(sort (mv_geq order) (mv_mmul(x,y) :: (mv_mul(p1,y::p2,order) @
neuper@37950
   911
									    mv_mul([x],p2,order)))),order);
neuper@37950
   912
neuper@37950
   913
(*. gets the maximum value of a list .*)
neuper@37950
   914
fun mv_getmax([])=0
neuper@37950
   915
  | mv_getmax(x::p1)= let 
neuper@37950
   916
		       val m=mv_getmax(p1);
neuper@37950
   917
		   in
neuper@37950
   918
		       if m>x then m
neuper@37950
   919
		       else x
neuper@37950
   920
		   end;
neuper@37950
   921
(*. calculates the maximum degree of an multivariate polynomial .*)
neuper@37950
   922
fun mv_grad([]:mv_poly) = 0 
neuper@37950
   923
  | mv_grad(p1:mv_poly)= mv_getmax((map mv_addlist) ((map #2) p1));
neuper@37950
   924
neuper@37950
   925
(*. converts the sign of a value .*)
neuper@37950
   926
fun mv_minus(x)=(~1) * x;
neuper@37950
   927
neuper@37950
   928
(*. converts the sign of all coefficients of a polynomial .*)
neuper@37950
   929
fun mv_minus2([]:mv_poly)=[]:mv_poly
neuper@37950
   930
  | mv_minus2(p1)=(mv_minus(#1(hd(p1))),#2(hd(p1)))::(mv_minus2(tl(p1)));
neuper@37950
   931
neuper@37950
   932
(*. searches for a negativ value in a list .*)
neuper@37950
   933
fun mv_is_negativ([])=false
neuper@37950
   934
  | mv_is_negativ(x::xs)=if x<0 then true else mv_is_negativ(xs);
neuper@37950
   935
neuper@37950
   936
(*. division of monomials .*)
neuper@37950
   937
fun mv_mdiv((0,[]):mv_monom,_:mv_monom)=(0,[]):mv_monom
neuper@38031
   938
  | mv_mdiv(_,(0,[]))= error ("RATIONALS_MV_MDIV_EXCEPTION Division by 0 ")
neuper@37950
   939
  | mv_mdiv(p1:mv_monom,p2:mv_monom)= 
neuper@37950
   940
    let
neuper@38006
   941
	val c= Unsynchronized.ref  (#1(p2));
neuper@38006
   942
	val pp= Unsynchronized.ref  [];
neuper@37950
   943
    in 
neuper@37950
   944
	(
neuper@38031
   945
	 if !c=0 then error("MV_MDIV_EXCEPTION Dividing by zero")
neuper@37950
   946
	 else c:=(#1(p1) div #1(p2));
neuper@37950
   947
	     if #1(p2)<>0 then 
neuper@37950
   948
		 (
neuper@37950
   949
		  pp:=(#2(mv_mmul((1,#2(p1)),(1,(map mv_minus) (#2(p2))))));
neuper@37950
   950
		  if mv_is_negativ(!pp) then (0,!pp)
neuper@37950
   951
		  else (!c,!pp) 
neuper@37950
   952
		      )
neuper@38031
   953
	     else error("MV_MDIV_EXCEPTION Dividing by empty Polynom")
neuper@37950
   954
		 )
neuper@37950
   955
    end;
neuper@37950
   956
neuper@37950
   957
(*. prints a polynom for (internal use only) .*)
neuper@38015
   958
fun mv_print_poly([]:mv_poly)=tracing("[]\n")
neuper@38015
   959
  | mv_print_poly((x,y)::[])= tracing("("^Int.toString(x)^","^ints2str(y)^")\n")
neuper@38015
   960
  | mv_print_poly((x,y)::p1) = (tracing("("^Int.toString(x)^","^ints2str(y)^"),");mv_print_poly(p1));
neuper@37950
   961
neuper@37950
   962
neuper@37950
   963
(*. division of two multivariate polynomials .*) 
neuper@37950
   964
fun mv_division([]:mv_poly,g:mv_poly,order)=([]:mv_poly,[]:mv_poly)
neuper@38031
   965
  | mv_division(f,[],order)= error ("RATIONALS_MV_DIVISION_EXCEPTION Division by zero")
neuper@37950
   966
  | mv_division(f,g,order)=
neuper@37950
   967
    let 
neuper@38006
   968
	val r= Unsynchronized.ref  [];
neuper@38006
   969
	val q= Unsynchronized.ref  [];
neuper@38006
   970
	val g'= Unsynchronized.ref  ([] : mv_monom list);
neuper@38006
   971
	val k= Unsynchronized.ref  0;
neuper@38006
   972
	val m= Unsynchronized.ref  (0,[0]);
neuper@38006
   973
	val exit= Unsynchronized.ref  0;
neuper@37950
   974
    in
neuper@37950
   975
	r := rev(sort (mv_geq order) (mv_shorten(f,order)));
neuper@37950
   976
	g':= rev(sort (mv_geq order) (mv_shorten(g,order)));
neuper@38031
   977
	if #1(hd(!g'))=0 then error("RATIONALS_MV_DIVISION_EXCEPTION: Dividing by zero") else ();
neuper@37950
   978
	if  (mv_monom_greater (hd(!g'),hd(!r),order)) then ([(0,mv_null2(#2(hd(f))))],(!r))
neuper@37950
   979
	else
neuper@37950
   980
	    (
neuper@37950
   981
	     exit:=0;
neuper@37950
   982
	     while (if (!exit)=0 then not(mv_monom_greater (hd(!g'),hd(!r),order)) else false) do
neuper@37950
   983
		 (
neuper@37950
   984
		  if (#1(mv_lm(!g',order)))<>0 then m:=mv_mdiv(mv_lm(!r,order),mv_lm(!g',order))
neuper@38031
   985
		  else error ("RATIONALS_MV_DIVISION_EXCEPTION: Dividing by zero");	  
neuper@37950
   986
		  if #1(!m)<>0 then
neuper@37950
   987
		      ( 
neuper@37950
   988
		       q:=(!m)::(!q);
neuper@37950
   989
		       r:=mv_add((!r),mv_minus2(mv_mul(!g',[!m],order)),order)
neuper@37950
   990
		       )
neuper@37950
   991
		  else exit:=1;
neuper@37950
   992
		  if (if length(!r)<>0 then length(!g')<>0 else false) then ()
neuper@37950
   993
		  else (exit:=1)
neuper@37950
   994
		  );
neuper@37950
   995
		 (rev(!q),!r)
neuper@37950
   996
		 )
neuper@37950
   997
    end;
neuper@37950
   998
neuper@37950
   999
(*. multiplies a polynomial with an integer .*)
neuper@37950
  1000
fun mv_skalar_mul([]:mv_poly,c) = []:mv_poly
neuper@37950
  1001
  | mv_skalar_mul((x,y)::p1,c) = ((x * c),y)::mv_skalar_mul(p1,c); 
neuper@37950
  1002
neuper@37950
  1003
(*. inserts the a first variable into an polynomial with exponent v .*)
neuper@37950
  1004
fun mv_correct([]:mv_poly,v:int)=[]:mv_poly
neuper@37950
  1005
  | mv_correct((x,y)::list,v:int)=(x,v::y)::mv_correct(list,v);
neuper@37950
  1006
neuper@37950
  1007
(*. multivariate case .*)
neuper@37950
  1008
neuper@37950
  1009
(*. decides if x is a factor of y .*)
neuper@38031
  1010
fun mv_divides([]:mv_poly,[]:mv_poly)=  error("RATIONALS_MV_DIVIDES_EXCEPTION: division by zero")
neuper@38031
  1011
  | mv_divides(x,[]) =  error("RATIONALS_MV_DIVIDES_EXCEPTION: division by zero")
neuper@37950
  1012
  | mv_divides(x:mv_poly,y:mv_poly) = #2(mv_division(y,x,LEX_))=[];
neuper@37950
  1013
neuper@37950
  1014
(*. gets the maximum of a and b .*)
neuper@37950
  1015
fun mv_max(a,b) = if a>b then a else b;
neuper@37950
  1016
neuper@37950
  1017
(*. gets the maximum exponent of a mv polynomial in the lexicographic term order .*)
neuper@37950
  1018
fun mv_deg([]:mv_poly) = 0  
neuper@37950
  1019
  | mv_deg(p1)=
neuper@37950
  1020
    let
neuper@37950
  1021
	val p1'=mv_shorten(p1,LEX_);
neuper@37950
  1022
    in
neuper@37950
  1023
	if length(p1')=0 then 0 
neuper@37950
  1024
	else mv_max(hd(#2(hd(p1'))),mv_deg(tl(p1')))
neuper@37950
  1025
    end;
neuper@37950
  1026
neuper@37950
  1027
(*. gets the maximum exponent of a mv polynomial in the reverse lexicographic term order .*)
neuper@37950
  1028
fun mv_deg2([]:mv_poly) = 0
neuper@37950
  1029
  | mv_deg2(p1)=
neuper@37950
  1030
    let
neuper@37950
  1031
	val p1'=mv_shorten(p1,LEX_);
neuper@37950
  1032
    in
neuper@37950
  1033
	if length(p1')=0 then 0 
neuper@37950
  1034
	else mv_max(hd(rev(#2(hd(p1')))),mv_deg2(tl(p1')))
neuper@37950
  1035
    end;
neuper@37950
  1036
neuper@37950
  1037
(*. evaluates the mv polynomial at the value v of the main variable .*)
neuper@37950
  1038
fun mv_subs([]:mv_poly,v) = []:mv_poly
neuper@37950
  1039
  | mv_subs((c,e)::p1:mv_poly,v) = mv_skalar_mul(mv_cut([(c,e)]),power v (hd(e))) @ mv_subs(p1,v);
neuper@37950
  1040
neuper@37950
  1041
(*. calculates the content of a uv-polynomial in mv-representation .*)
neuper@37950
  1042
fun uv_content2([]:mv_poly) = 0
neuper@37950
  1043
  | uv_content2((c,e)::p1) = (gcd_int c (uv_content2(p1)));
neuper@37950
  1044
neuper@37950
  1045
(*. converts a uv-polynomial from mv-representation to  uv-representation .*)
neuper@37950
  1046
fun uv_to_list ([]:mv_poly)=[]:uv_poly
neuper@37950
  1047
  | uv_to_list ((c1,e1)::others) = 
neuper@37950
  1048
    let
neuper@38006
  1049
	val count= Unsynchronized.ref  0;
neuper@37950
  1050
	val max=mv_grad((c1,e1)::others); 
neuper@38006
  1051
	val help= Unsynchronized.ref  ((c1,e1)::others);
neuper@38006
  1052
	val list= Unsynchronized.ref  [];
neuper@37950
  1053
    in
neuper@38031
  1054
	if length(e1)>1 then error ("RATIONALS_TO_LIST_EXCEPTION: not univariate")
neuper@37950
  1055
	else if length(e1)=0 then [c1]
neuper@37950
  1056
	     else
neuper@37950
  1057
		 (
neuper@37950
  1058
		  count:=0;
neuper@37950
  1059
		  while (!count)<=max do
neuper@37950
  1060
		      (
neuper@37950
  1061
		       if length(!help)>0 andalso hd(#2(hd(!help)))=max-(!count) then 
neuper@37950
  1062
			   (
neuper@37950
  1063
			    list:=(#1(hd(!help)))::(!list);		       
neuper@37950
  1064
			    help:=tl(!help) 
neuper@37950
  1065
			    )
neuper@37950
  1066
		       else 
neuper@37950
  1067
			   (
neuper@37950
  1068
			    list:= 0::(!list)
neuper@37950
  1069
			    );
neuper@37950
  1070
		       count := (!count) + 1
neuper@37950
  1071
		       );
neuper@37950
  1072
		      (!list)
neuper@37950
  1073
		      )
neuper@37950
  1074
    end;
neuper@37950
  1075
neuper@37950
  1076
(*. converts a uv-polynomial from uv-representation to mv-representation .*)
neuper@37950
  1077
fun uv_to_poly ([]:uv_poly) = []:mv_poly
neuper@37950
  1078
  | uv_to_poly p1 = 
neuper@37950
  1079
    let
neuper@38006
  1080
	val count= Unsynchronized.ref  0;
neuper@38006
  1081
	val help= Unsynchronized.ref  p1;
neuper@38006
  1082
	val list= Unsynchronized.ref  [];
neuper@37950
  1083
    in
neuper@37950
  1084
	while length(!help)>0 do
neuper@37950
  1085
	    (
neuper@37950
  1086
	     if hd(!help)=0 then ()
neuper@37950
  1087
	     else list:=(hd(!help),[!count])::(!list);
neuper@37950
  1088
	     count:=(!count)+1;
neuper@37950
  1089
	     help:=tl(!help)
neuper@37950
  1090
	     );
neuper@37950
  1091
	    (!list)
neuper@37950
  1092
    end;
neuper@37950
  1093
neuper@37950
  1094
(*. univariate gcd calculation from polynomials in multivariate representation .*)
neuper@37950
  1095
fun uv_gcd ([]:mv_poly) p2 = p2
neuper@37950
  1096
  | uv_gcd p1 ([]:mv_poly) = p1
neuper@37950
  1097
  | uv_gcd p1 [(c,[e])] = 
neuper@37950
  1098
    let 
neuper@38006
  1099
	val list= Unsynchronized.ref  (rev(sort (mv_geq LEX_) (mv_shorten(p1,LEX_))));
neuper@37950
  1100
	val min=uv_mod_min(e,(hd(#2(hd(rev(!list))))));
neuper@37950
  1101
    in
neuper@37950
  1102
	[(gcd_int (uv_content2(p1)) c,[min])]
neuper@37950
  1103
    end
neuper@37950
  1104
  | uv_gcd [(c,[e])] p2 = 
neuper@37950
  1105
    let 
neuper@38006
  1106
	val list= Unsynchronized.ref  (rev(sort (mv_geq LEX_) (mv_shorten(p2,LEX_))));
neuper@37950
  1107
	val min=uv_mod_min(e,(hd(#2(hd(rev(!list))))));
neuper@37950
  1108
    in
neuper@37950
  1109
	[(gcd_int (uv_content2(p2)) c,[min])]
neuper@37950
  1110
    end 
neuper@37950
  1111
  | uv_gcd p11 p22 = uv_to_poly(uv_mod_gcd (uv_to_list(mv_shorten(p11,LEX_))) (uv_to_list(mv_shorten(p22,LEX_))));
neuper@37950
  1112
neuper@37950
  1113
(*. help function for the newton interpolation .*)
neuper@37950
  1114
fun mv_newton_help ([]:mv_poly list,k:int) = []:mv_poly list
neuper@37950
  1115
  | mv_newton_help (pl:mv_poly list,k) = 
neuper@37950
  1116
    let
neuper@38006
  1117
	val x= Unsynchronized.ref  (rev(pl));
neuper@38006
  1118
	val t= Unsynchronized.ref  [];
neuper@38006
  1119
	val y= Unsynchronized.ref  [];
neuper@38006
  1120
	val n= Unsynchronized.ref  1;
neuper@38006
  1121
	val n1= Unsynchronized.ref [];
neuper@37950
  1122
    in
neuper@37950
  1123
	(
neuper@37950
  1124
	 while length(!x)>1 do 
neuper@37950
  1125
	     (
neuper@37950
  1126
	      if length(hd(!x))>0 then n1:=mv_null2(#2(hd(hd(!x))))
neuper@37950
  1127
	      else if length(hd(tl(!x)))>0 then n1:=mv_null2(#2(hd(hd(tl(!x)))))
neuper@37950
  1128
		   else n1:=[]; 
neuper@37950
  1129
	      t:= #1(mv_division(mv_add(hd(!x),mv_skalar_mul(hd(tl(!x)),~1),LEX_),[(k,!n1)],LEX_)); 
neuper@37950
  1130
	      y:=(!t)::(!y);
neuper@37950
  1131
	      x:=tl(!x)
neuper@37950
  1132
	      );
neuper@37950
  1133
	 (!y)
neuper@37950
  1134
	 )
neuper@37950
  1135
    end;
neuper@37950
  1136
    
neuper@37950
  1137
(*. help function for the newton interpolation .*)
neuper@37950
  1138
fun mv_newton_add ([]:mv_poly list) t = []:mv_poly
neuper@37950
  1139
  | mv_newton_add [x:mv_poly] t = x 
neuper@37950
  1140
  | mv_newton_add (pl:mv_poly list) t = 
neuper@37950
  1141
    let
neuper@38006
  1142
	val expos= Unsynchronized.ref  [];
neuper@38006
  1143
	val pll= Unsynchronized.ref  pl;
neuper@37950
  1144
    in
neuper@37950
  1145
	(
neuper@37950
  1146
neuper@37950
  1147
	 while length(!pll)>0 andalso hd(!pll)=[]  do 
neuper@37950
  1148
	     ( 
neuper@37950
  1149
	      pll:=tl(!pll)
neuper@37950
  1150
	      ); 
neuper@37950
  1151
	 if length(!pll)>0 then expos:= #2(hd(hd(!pll))) else expos:=[];
neuper@37950
  1152
	 mv_add(hd(pl),
neuper@37950
  1153
		mv_mul(
neuper@37950
  1154
		       mv_add(mv_correct(mv_cut([(1,mv_null2(!expos))]),1),[(~t,mv_null2(!expos))],LEX_),
neuper@37950
  1155
		       mv_newton_add (tl(pl)) (t+1),
neuper@37950
  1156
		       LEX_
neuper@37950
  1157
		       ),
neuper@37950
  1158
		LEX_)
neuper@37950
  1159
	 )
neuper@37950
  1160
    end;
neuper@37950
  1161
neuper@37950
  1162
(*. calculates the newton interpolation with polynomial coefficients .*)
neuper@37950
  1163
(*. step-depth is 1 and if the result is not an integerpolynomial .*)
neuper@37950
  1164
(*. this function returns [] .*)
neuper@37950
  1165
fun mv_newton ([]:(mv_poly) list) = []:mv_poly 
neuper@37950
  1166
  | mv_newton ([mp]:(mv_poly) list) = mp:mv_poly
neuper@37950
  1167
  | mv_newton pl =
neuper@37950
  1168
    let
neuper@38006
  1169
	val c= Unsynchronized.ref  pl;
neuper@38006
  1170
	val c1= Unsynchronized.ref  [];
neuper@37950
  1171
	val n=length(pl);
neuper@38006
  1172
	val k= Unsynchronized.ref  1;
neuper@38006
  1173
	val i= Unsynchronized.ref  n;
neuper@38006
  1174
	val ppl= Unsynchronized.ref  [];
neuper@37950
  1175
    in
neuper@37950
  1176
	c1:=hd(pl)::[];
neuper@37950
  1177
	c:=mv_newton_help(!c,!k);
neuper@37950
  1178
	c1:=(hd(!c))::(!c1);
neuper@37950
  1179
	while(length(!c)>1 andalso !k<n) do
neuper@37950
  1180
	    (	 
neuper@37950
  1181
	     k:=(!k)+1; 
neuper@37950
  1182
	     while  length(!c)>0 andalso hd(!c)=[] do c:=tl(!c); 	  
neuper@37950
  1183
	     if !c=[] then () else c:=mv_newton_help(!c,!k);
neuper@37950
  1184
	     ppl:= !c;
neuper@37950
  1185
	     if !c=[] then () else  c1:=(hd(!c))::(!c1)
neuper@37950
  1186
	     );
neuper@37950
  1187
	while  hd(!c1)=[] do c1:=tl(!c1);
neuper@37950
  1188
	c1:=rev(!c1);
neuper@37950
  1189
	ppl:= !c1;
neuper@37950
  1190
	mv_newton_add (!c1) 1
neuper@37950
  1191
    end;
neuper@37950
  1192
neuper@37950
  1193
(*. sets the exponents of the first variable to zero .*)
neuper@37950
  1194
fun mv_null3([]:mv_poly)    = []:mv_poly
neuper@37950
  1195
  | mv_null3((x,y)::xs) = (x,0::tl(y))::mv_null3(xs);
neuper@37950
  1196
neuper@37950
  1197
(*. calculates the minimum exponents of a multivariate polynomial .*)
neuper@37950
  1198
fun mv_min_pp([]:mv_poly)=[]
neuper@37950
  1199
  | mv_min_pp((c,e)::xs)=
neuper@37950
  1200
    let
neuper@38006
  1201
	val y= Unsynchronized.ref  xs;
neuper@38006
  1202
	val x= Unsynchronized.ref  [];
neuper@37950
  1203
    in
neuper@37950
  1204
	(
neuper@37950
  1205
	 x:=e;
neuper@37950
  1206
	 while length(!y)>0 do
neuper@37950
  1207
	     (
neuper@37950
  1208
	      x:=(map uv_mod_min) ((!x) ~~ (#2(hd(!y))));
neuper@37950
  1209
	      y:=tl(!y)
neuper@37950
  1210
	      );
neuper@37950
  1211
	 !x
neuper@37950
  1212
	 )
neuper@37950
  1213
    end;
neuper@37950
  1214
neuper@37950
  1215
(*. checks if all elements of the list have value zero .*)
neuper@37950
  1216
fun list_is_null [] = true 
neuper@37950
  1217
  | list_is_null (x::xs) = (x=0 andalso list_is_null(xs)); 
neuper@37950
  1218
neuper@37950
  1219
(* check if main variable is zero*)
neuper@37950
  1220
fun main_zero (ms : mv_poly) = (list_is_null o (map (hd o #2))) ms;
neuper@37950
  1221
neuper@37950
  1222
(*. calculates the content of an polynomial .*)
neuper@37950
  1223
fun mv_content([]:mv_poly) = []:mv_poly
neuper@37950
  1224
  | mv_content(p1) = 
neuper@37950
  1225
    let
neuper@38006
  1226
	val list= Unsynchronized.ref  (rev(sort (mv_geq LEX_) (mv_shorten(p1,LEX_))));
neuper@38006
  1227
	val test= Unsynchronized.ref  (hd(#2(hd(!list))));
neuper@38006
  1228
	val result= Unsynchronized.ref  []; 
neuper@37950
  1229
	val min=(hd(#2(hd(rev(!list)))));
neuper@37950
  1230
    in
neuper@37950
  1231
	(
neuper@37950
  1232
	 if length(!list)>1 then
neuper@37950
  1233
	     (
neuper@37950
  1234
	      while (if length(!list)>0 then (hd(#2(hd(!list)))=(!test)) else false) do
neuper@37950
  1235
		  (
neuper@37950
  1236
		   result:=(#1(hd(!list)),tl(#2(hd(!list))))::(!result);
neuper@37950
  1237
		   
neuper@37950
  1238
		   if length(!list)<1 then list:=[]
neuper@37950
  1239
		   else list:=tl(!list) 
neuper@37950
  1240
		       
neuper@37950
  1241
		       );		  
neuper@37950
  1242
		  if length(!list)>0 then  
neuper@37950
  1243
		   ( 
neuper@37950
  1244
		    list:=mv_gcd (!result) (mv_cut(mv_content(!list))) 
neuper@37950
  1245
		    ) 
neuper@37950
  1246
		  else list:=(!result); 
neuper@37950
  1247
		  list:=mv_correct(!list,0);  
neuper@37950
  1248
		  (!list) 
neuper@37950
  1249
		  )
neuper@37950
  1250
	 else
neuper@37950
  1251
	     (
neuper@37950
  1252
	      mv_null3(!list) 
neuper@37950
  1253
	      )
neuper@37950
  1254
	     )
neuper@37950
  1255
    end
neuper@37950
  1256
neuper@37950
  1257
(*. calculates the primitiv part of a polynomial .*)
neuper@37950
  1258
and mv_pp([]:mv_poly) = []:mv_poly
neuper@37950
  1259
  | mv_pp(p1) = let
neuper@38006
  1260
		    val cont= Unsynchronized.ref  []; 
neuper@38006
  1261
		    val pp= Unsynchronized.ref [];
neuper@37950
  1262
		in
neuper@37950
  1263
		    cont:=mv_content(p1);
neuper@37950
  1264
		    pp:=(#1(mv_division(p1,!cont,LEX_)));
neuper@37950
  1265
		    if !pp=[] 
neuper@38031
  1266
			then error("RATIONALS_MV_PP_EXCEPTION: Invalid Content ")
neuper@37950
  1267
		    else (!pp)
neuper@37950
  1268
		end
neuper@37950
  1269
neuper@37950
  1270
(*. calculates the gcd of two multivariate polynomials with a modular approach .*)
neuper@37950
  1271
and mv_gcd ([]:mv_poly) ([]:mv_poly) :mv_poly= []:mv_poly
neuper@37950
  1272
  | mv_gcd ([]:mv_poly) (p2) :mv_poly= p2:mv_poly
neuper@37950
  1273
  | mv_gcd (p1:mv_poly) ([]) :mv_poly= p1:mv_poly
neuper@37950
  1274
  | mv_gcd ([(x,xs)]:mv_poly) ([(y,ys)]):mv_poly = 
neuper@37950
  1275
     let
neuper@37950
  1276
      val xpoly:mv_poly = [(x,xs)];
neuper@37950
  1277
      val ypoly:mv_poly = [(y,ys)];
neuper@37950
  1278
     in 
neuper@37950
  1279
	(
neuper@37950
  1280
	 if xs=ys then [((gcd_int x y),xs)]
neuper@37950
  1281
	 else [((gcd_int x y),(map uv_mod_min)(xs~~ys))]:mv_poly
neuper@37950
  1282
        )
neuper@37950
  1283
    end 
neuper@37950
  1284
  | mv_gcd (p1:mv_poly) ([(y,ys)]) :mv_poly= 
neuper@37950
  1285
	(
neuper@37950
  1286
	 [(gcd_int (uv_content2(p1)) (y),(map  uv_mod_min)(mv_min_pp(p1)~~ys))]:mv_poly
neuper@37950
  1287
	)
neuper@37950
  1288
  | mv_gcd ([(y,ys)]:mv_poly) (p2):mv_poly = 
neuper@37950
  1289
	(
neuper@37950
  1290
         [(gcd_int (uv_content2(p2)) (y),(map  uv_mod_min)(mv_min_pp(p2)~~ys))]:mv_poly
neuper@37950
  1291
        )
neuper@37950
  1292
  | mv_gcd (p1':mv_poly) (p2':mv_poly):mv_poly=
neuper@37950
  1293
    let
neuper@37950
  1294
	val vc=length(#2(hd(p1')));
neuper@37950
  1295
	val cont = 
neuper@37950
  1296
		  (
neuper@37950
  1297
                   if main_zero(mv_content(p1')) andalso 
neuper@37950
  1298
                     (main_zero(mv_content(p2'))) then
neuper@37950
  1299
                     mv_correct((mv_gcd (mv_cut(mv_content(p1'))) (mv_cut(mv_content(p2')))),0)
neuper@37950
  1300
                   else 
neuper@37950
  1301
                     mv_gcd (mv_content(p1')) (mv_content(p2'))
neuper@37950
  1302
                  );
neuper@37950
  1303
	val p1= #1(mv_division(p1',mv_content(p1'),LEX_));
neuper@37950
  1304
	val p2= #1(mv_division(p2',mv_content(p2'),LEX_)); 
neuper@38006
  1305
	val gcd= Unsynchronized.ref  [];
neuper@38006
  1306
	val candidate= Unsynchronized.ref  [];
neuper@38006
  1307
	val interpolation_list= Unsynchronized.ref  [];
neuper@38006
  1308
	val delta= Unsynchronized.ref  [];
neuper@38006
  1309
        val p1r = Unsynchronized.ref [];
neuper@38006
  1310
        val p2r = Unsynchronized.ref [];
neuper@38006
  1311
        val p1r' = Unsynchronized.ref [];
neuper@38006
  1312
        val p2r' = Unsynchronized.ref [];
neuper@38006
  1313
	val factor= Unsynchronized.ref  [];
neuper@38006
  1314
	val r= Unsynchronized.ref  0;
neuper@38006
  1315
	val gcd_r= Unsynchronized.ref  [];
neuper@38006
  1316
	val d= Unsynchronized.ref  0;
neuper@38006
  1317
	val exit= Unsynchronized.ref  0;
neuper@38006
  1318
	val current_degree= Unsynchronized.ref  99999; (*. FIXME: unlimited ! .*)
neuper@37950
  1319
    in
neuper@37950
  1320
	(
neuper@37950
  1321
	 if vc<2 then (* areUnivariate(p1',p2') *)
neuper@37950
  1322
	     (
neuper@37950
  1323
	      gcd:=uv_gcd (mv_shorten(p1',LEX_)) (mv_shorten(p2',LEX_))
neuper@37950
  1324
	      )
neuper@37950
  1325
	 else
neuper@37950
  1326
	     (
neuper@37950
  1327
	      while !exit=0 do
neuper@37950
  1328
		  (
neuper@37950
  1329
		   r:=(!r)+1;
neuper@37950
  1330
                   p1r := mv_lc(p1,LEX_);
neuper@37950
  1331
		   p2r := mv_lc(p2,LEX_);
neuper@37950
  1332
                   if main_zero(!p1r) andalso
neuper@37950
  1333
                      main_zero(!p2r) 
neuper@37950
  1334
                   then
neuper@37950
  1335
                       (
neuper@37950
  1336
                        delta := mv_correct((mv_gcd (mv_cut (!p1r)) (mv_cut (!p2r))),0)
neuper@37950
  1337
                       )
neuper@37950
  1338
                   else
neuper@37950
  1339
                       (
neuper@37950
  1340
		        delta := mv_gcd (!p1r) (!p2r)
neuper@37950
  1341
                       );
neuper@37950
  1342
		   (*if mv_shorten(mv_subs(!p1r,!r),LEX_)=[] andalso 
neuper@37950
  1343
		      mv_shorten(mv_subs(!p2r,!r),LEX_)=[] *)
neuper@37950
  1344
		   if mv_lc2(mv_shorten(mv_subs(!p1r,!r),LEX_),LEX_)=0 andalso 
neuper@37950
  1345
		      mv_lc2(mv_shorten(mv_subs(!p2r,!r),LEX_),LEX_)=0 
neuper@37950
  1346
                   then 
neuper@37950
  1347
                       (
neuper@37950
  1348
		       )
neuper@37950
  1349
		   else 
neuper@37950
  1350
		       (
neuper@37950
  1351
			gcd_r:=mv_shorten(mv_gcd (mv_shorten(mv_subs(p1,!r),LEX_)) 
neuper@37950
  1352
					         (mv_shorten(mv_subs(p2,!r),LEX_)) ,LEX_);
neuper@37950
  1353
			gcd_r:= #1(mv_division(mv_mul(mv_correct(mv_subs(!delta,!r),0),!gcd_r,LEX_),
neuper@37950
  1354
					       mv_correct(mv_lc(!gcd_r,LEX_),0),LEX_));
neuper@37950
  1355
			d:=mv_deg2(!gcd_r); (* deg(gcd_r,z) *)
neuper@37950
  1356
			if (!d < !current_degree) then 
neuper@37950
  1357
			    (
neuper@37950
  1358
			     current_degree:= !d;
neuper@37950
  1359
			     interpolation_list:=mv_correct(!gcd_r,0)::(!interpolation_list)
neuper@37950
  1360
			     )
neuper@37950
  1361
			else
neuper@37950
  1362
			    (
neuper@37950
  1363
			     if (!d = !current_degree) then
neuper@37950
  1364
				 (
neuper@37950
  1365
				  interpolation_list:=mv_correct(!gcd_r,0)::(!interpolation_list)
neuper@37950
  1366
				  )
neuper@37950
  1367
			     else () 
neuper@37950
  1368
				 )
neuper@37950
  1369
			    );
neuper@37950
  1370
		      if length(!interpolation_list)> uv_mod_min(mv_deg(p1),mv_deg(p2)) then 
neuper@37950
  1371
			  (
neuper@37950
  1372
			   candidate := mv_newton(rev(!interpolation_list));
neuper@37950
  1373
			   if !candidate=[] then ()
neuper@37950
  1374
			   else
neuper@37950
  1375
			       (
neuper@37950
  1376
				candidate:=mv_pp(!candidate);
neuper@37950
  1377
				if mv_divides(!candidate,p1) andalso mv_divides(!candidate,p2) then
neuper@37950
  1378
				    (
neuper@37950
  1379
				     gcd:= mv_mul(!candidate,cont,LEX_);
neuper@37950
  1380
				     exit:=1
neuper@37950
  1381
				     )
neuper@37950
  1382
				else ()
neuper@37950
  1383
				    );
neuper@37950
  1384
			       interpolation_list:=[mv_correct(!gcd_r,0)]
neuper@37950
  1385
			       )
neuper@37950
  1386
		      else ()
neuper@37950
  1387
			  )
neuper@37950
  1388
	     );
neuper@37950
  1389
	     (!gcd):mv_poly
neuper@37950
  1390
	     )
neuper@37950
  1391
    end;	
neuper@37950
  1392
neuper@37950
  1393
neuper@37950
  1394
(*. calculates the least common divisor of two polynomials .*)
neuper@37950
  1395
fun mv_lcm (p1:mv_poly) (p2:mv_poly) :mv_poly = 
neuper@37950
  1396
    (
neuper@37950
  1397
     #1(mv_division(mv_mul(p1,p2,LEX_),mv_gcd p1 p2,LEX_))
neuper@37950
  1398
     );
neuper@37950
  1399
neuper@42391
  1400
(* gets the variables (strings) of a term *)
neuper@42391
  1401
neuper@37950
  1402
fun get_vars(term1) = (map free2str) (vars term1); (*["a","b","c"]; *)
neuper@37950
  1403
neuper@37950
  1404
(*. counts the negative coefficents in a polynomial .*)
neuper@37950
  1405
fun count_neg ([]:mv_poly) = 0 
neuper@37950
  1406
  | count_neg ((c,e)::xs) = if c<0 then 1+count_neg xs
neuper@37950
  1407
			  else count_neg xs;
neuper@37950
  1408
neuper@37950
  1409
(*. help function for is_polynomial  
neuper@37950
  1410
    checks the order of the operators .*)
neuper@37950
  1411
fun test_polynomial (Const ("uminus",_) $ Free (str,_)) _ = true (*WN.13.3.03*)
neuper@37950
  1412
  | test_polynomial (t as Free(str,_)) v = true
neuper@38034
  1413
  | test_polynomial (t as Const ("Groups.times_class.times",_) $ t1 $ t2) v = if v="^" then false
neuper@37950
  1414
						     else (test_polynomial t1 "*") andalso (test_polynomial t2 "*")
neuper@38014
  1415
  | test_polynomial (t as Const ("Groups.plus_class.plus",_) $ t1 $ t2) v = if v="*" orelse v="^" then false 
neuper@37950
  1416
							  else (test_polynomial t1 " ") andalso (test_polynomial t2 " ")
neuper@37950
  1417
  | test_polynomial (t as Const ("Atools.pow",_) $ t1 $ t2) v = (test_polynomial t1 "^") andalso (test_polynomial t2 "^")
neuper@37950
  1418
  | test_polynomial _ v = false;  
neuper@37950
  1419
neuper@37950
  1420
(*. tests if a term is a polynomial .*)  
neuper@37950
  1421
fun is_polynomial t = test_polynomial t " ";
neuper@37950
  1422
neuper@37950
  1423
(*. help function for is_expanded 
neuper@37950
  1424
    checks the order of the operators .*)
neuper@37950
  1425
fun test_exp (t as Free(str,_)) v = true 
neuper@38034
  1426
  | test_exp (t as Const ("Groups.times_class.times",_) $ t1 $ t2) v = if v="^" then false
neuper@37950
  1427
						     else (test_exp t1 "*") andalso (test_exp t2 "*")
neuper@38014
  1428
  | test_exp (t as Const ("Groups.plus_class.plus",_) $ t1 $ t2) v = if v="*" orelse v="^" then false 
neuper@37950
  1429
							  else (test_exp t1 " ") andalso (test_exp t2 " ") 
neuper@38014
  1430
  | test_exp (t as Const ("Groups.minus_class.minus",_) $ t1 $ t2) v = if v="*" orelse v="^" then false 
neuper@37950
  1431
							  else (test_exp t1 " ") andalso (test_exp t2 " ")
neuper@37950
  1432
  | test_exp (t as Const ("Atools.pow",_) $ t1 $ t2) v = (test_exp t1 "^") andalso (test_exp t2 "^")
neuper@37950
  1433
  | test_exp  _ v = false;
neuper@37950
  1434
neuper@37950
  1435
neuper@37950
  1436
(*. help function for check_coeff: 
neuper@37950
  1437
    converts the term to a list of coefficients .*) 
neuper@37950
  1438
fun term2coef' (t as Free(str,_(*typ*))) v :mv_poly option = 
neuper@37950
  1439
    let
neuper@38006
  1440
	val x= Unsynchronized.ref  NONE;
neuper@38006
  1441
	val len= Unsynchronized.ref  0;
neuper@38006
  1442
	val vl= Unsynchronized.ref  [];
neuper@38006
  1443
	val vh= Unsynchronized.ref  [];
neuper@38006
  1444
	val i= Unsynchronized.ref  0;
neuper@37950
  1445
    in 
neuper@37950
  1446
	if is_numeral str then
neuper@37950
  1447
	    (
neuper@37950
  1448
	     SOME [(((the o int_of_str) str),mv_null2(v))] handle _ => NONE
neuper@37950
  1449
		 )
neuper@37950
  1450
	else (* variable *)
neuper@37950
  1451
	    (
neuper@37950
  1452
	     len:=length(v);
neuper@37950
  1453
	     vh:=v;
neuper@37950
  1454
	     while ((!len)>(!i)) do
neuper@37950
  1455
		 (
neuper@37950
  1456
		  if str=hd((!vh)) then
neuper@37950
  1457
		      (
neuper@37950
  1458
		       vl:=1::(!vl)
neuper@37950
  1459
		       )
neuper@37950
  1460
		  else 
neuper@37950
  1461
		      (
neuper@37950
  1462
		       vl:=0::(!vl)
neuper@37950
  1463
		       );
neuper@37950
  1464
		      vh:=tl(!vh);
neuper@37950
  1465
		      i:=(!i)+1    
neuper@37950
  1466
		      );		
neuper@37950
  1467
		 SOME [(1,rev(!vl))] handle _ => NONE
neuper@37950
  1468
	    )
neuper@37950
  1469
    end
neuper@38034
  1470
  | term2coef' (Const ("Groups.times_class.times",_) $ t1 $ t2) v :mv_poly option= 
neuper@37950
  1471
    let
neuper@38006
  1472
	val t1pp= Unsynchronized.ref  [];
neuper@38006
  1473
	val t2pp= Unsynchronized.ref  [];
neuper@38006
  1474
	val t1c= Unsynchronized.ref  0;
neuper@38006
  1475
	val t2c= Unsynchronized.ref  0;
neuper@37950
  1476
    in
neuper@37950
  1477
	(
neuper@37950
  1478
	 t1pp:=(#2(hd(the(term2coef' t1 v))));
neuper@37950
  1479
	 t2pp:=(#2(hd(the(term2coef' t2 v))));
neuper@37950
  1480
	 t1c:=(#1(hd(the(term2coef' t1 v))));
neuper@37950
  1481
	 t2c:=(#1(hd(the(term2coef' t2 v))));
neuper@37950
  1482
	
neuper@37950
  1483
	 SOME [( (!t1c)*(!t2c) ,( (map op+) ((!t1pp)~~(!t2pp)) ) )] handle _ => NONE 
neuper@37950
  1484
		
neuper@37950
  1485
	 )
neuper@37950
  1486
    end
neuper@37950
  1487
  | term2coef' (Const ("Atools.pow",_) $ (t1 as Free (str1,_)) $ (t2 as Free (str2,_))) v :mv_poly option= 
neuper@37950
  1488
    let
neuper@38006
  1489
	val x= Unsynchronized.ref  NONE;
neuper@38006
  1490
	val len= Unsynchronized.ref  0;
neuper@38006
  1491
	val vl= Unsynchronized.ref  [];
neuper@38006
  1492
	val vh= Unsynchronized.ref  [];
neuper@38006
  1493
	val vtemp= Unsynchronized.ref  [];
neuper@38006
  1494
	val i= Unsynchronized.ref  0;	 
neuper@37950
  1495
    in
neuper@37950
  1496
    (
neuper@37950
  1497
     if (not o is_numeral) str1 andalso is_numeral str2 then
neuper@37950
  1498
	 (
neuper@37950
  1499
	  len:=length(v);
neuper@37950
  1500
	  vh:=v;
neuper@37950
  1501
neuper@37950
  1502
	  while ((!len)>(!i)) do
neuper@37950
  1503
	      (
neuper@37950
  1504
	       if str1=hd((!vh)) then
neuper@37950
  1505
		   (
neuper@37950
  1506
		    vl:=((the o int_of_str) str2)::(!vl)
neuper@37950
  1507
		    )
neuper@37950
  1508
	       else 
neuper@37950
  1509
		   (
neuper@37950
  1510
		    vl:=0::(!vl)
neuper@37950
  1511
		    );
neuper@37950
  1512
		   vh:=tl(!vh);
neuper@37950
  1513
		   i:=(!i)+1     
neuper@37950
  1514
		   );
neuper@37950
  1515
	      SOME [(1,rev(!vl))] handle _ => NONE
neuper@37950
  1516
	      )
neuper@38031
  1517
     else error ("RATIONALS_TERM2COEF_EXCEPTION 1: Invalid term")
neuper@37950
  1518
	 )
neuper@37950
  1519
    end
neuper@38014
  1520
  | term2coef' (Const ("Groups.plus_class.plus",_) $ t1 $ t2) v :mv_poly option= 
neuper@37950
  1521
    (
neuper@37950
  1522
     SOME ((the(term2coef' t1 v)) @ (the(term2coef' t2 v))) handle _ => NONE
neuper@37950
  1523
	 )
neuper@38014
  1524
  | term2coef' (Const ("Groups.minus_class.minus",_) $ t1 $ t2) v :mv_poly option= 
neuper@37950
  1525
    (
neuper@37950
  1526
     SOME ((the(term2coef' t1 v)) @ mv_skalar_mul((the(term2coef' t2 v)),1)) handle _ => NONE
neuper@37950
  1527
	 )
neuper@38031
  1528
  | term2coef' (term) v = error ("RATIONALS_TERM2COEF_EXCEPTION 2: Invalid term");
neuper@37950
  1529
neuper@37950
  1530
(*. checks if all coefficients of a polynomial are positiv (except the first) .*)
neuper@37950
  1531
fun check_coeff t = (* erste Koeffizient kann <0 sein !!! *)
neuper@37950
  1532
    if count_neg(tl(the(term2coef' t (get_vars(t)))))=0 then true 
neuper@37950
  1533
    else false;
neuper@37950
  1534
neuper@37950
  1535
(*. checks for expanded term [3] .*)
neuper@37950
  1536
fun is_expanded t = test_exp t " " andalso check_coeff(t); 
neuper@37950
  1537
neuper@37950
  1538
(*WN.7.3.03 Hilfsfunktion f"ur term2poly'*)
neuper@37950
  1539
fun mk_monom v' p vs = 
neuper@37950
  1540
    let fun conv p (v: string) = if v'= v then p else 0
neuper@37950
  1541
    in map (conv p) vs end;
neuper@37950
  1542
(* mk_monom "y" 5 ["a","b","x","y","z"];
neuper@37950
  1543
val it = [0,0,0,5,0] : int list*)
neuper@37950
  1544
neuper@37950
  1545
(*. this function converts the term representation into the internal representation mv_poly .*)
neuper@37950
  1546
fun term2poly' (Const ("uminus",_) $ Free (str,_)) v = (*WN.7.3.03*)
neuper@37950
  1547
    if is_numeral str 
neuper@37950
  1548
    then SOME [((the o int_of_str) ("-"^str), mk_monom "#" 0 v)]
neuper@37950
  1549
    else SOME [(~1, mk_monom str 1 v)]
neuper@37950
  1550
neuper@37950
  1551
  | term2poly' (Free(str,_)) v :mv_poly option = 
neuper@37950
  1552
    let
neuper@38006
  1553
	val x= Unsynchronized.ref  NONE;
neuper@38006
  1554
	val len= Unsynchronized.ref  0;
neuper@38006
  1555
	val vl= Unsynchronized.ref  [];
neuper@38006
  1556
	val vh= Unsynchronized.ref  [];
neuper@38006
  1557
	val i= Unsynchronized.ref  0;
neuper@37950
  1558
    in 
neuper@37950
  1559
	if is_numeral str then
neuper@37950
  1560
	    (
neuper@37950
  1561
	     SOME [(((the o int_of_str) str),mv_null2 v)] handle _ => NONE
neuper@37950
  1562
		 )
neuper@37950
  1563
	else (* variable *)
neuper@37950
  1564
	    (
neuper@37950
  1565
	     len:=length v;
neuper@37950
  1566
	     vh:= v;
neuper@37950
  1567
	     while ((!len)>(!i)) do
neuper@37950
  1568
		 (
neuper@37950
  1569
		  if str=hd((!vh)) then
neuper@37950
  1570
		      (
neuper@37950
  1571
		       vl:=1::(!vl)
neuper@37950
  1572
		       )
neuper@37950
  1573
		  else 
neuper@37950
  1574
		      (
neuper@37950
  1575
		       vl:=0::(!vl)
neuper@37950
  1576
		       );
neuper@37950
  1577
		      vh:=tl(!vh);
neuper@37950
  1578
		      i:=(!i)+1    
neuper@37950
  1579
		      );		
neuper@37950
  1580
		 SOME [(1,rev(!vl))] handle _ => NONE
neuper@37950
  1581
	    )
neuper@37950
  1582
    end
neuper@38034
  1583
  | term2poly' (Const ("Groups.times_class.times",_) $ t1 $ t2) v :mv_poly option= 
neuper@37950
  1584
    let
neuper@38006
  1585
	val t1pp= Unsynchronized.ref  [];
neuper@38006
  1586
	val t2pp= Unsynchronized.ref  [];
neuper@38006
  1587
	val t1c= Unsynchronized.ref  0;
neuper@38006
  1588
	val t2c= Unsynchronized.ref  0;
neuper@37950
  1589
    in
neuper@37950
  1590
	(
neuper@37950
  1591
	 t1pp:=(#2(hd(the(term2poly' t1 v))));
neuper@37950
  1592
	 t2pp:=(#2(hd(the(term2poly' t2 v))));
neuper@37950
  1593
	 t1c:=(#1(hd(the(term2poly' t1 v))));
neuper@37950
  1594
	 t2c:=(#1(hd(the(term2poly' t2 v))));
neuper@37950
  1595
	
neuper@37950
  1596
	 SOME [( (!t1c)*(!t2c) ,( (map op+) ((!t1pp)~~(!t2pp)) ) )] 
neuper@37950
  1597
	 handle _ => NONE 
neuper@37950
  1598
		
neuper@37950
  1599
	 )
neuper@37950
  1600
    end
neuper@37950
  1601
  | term2poly' (Const ("Atools.pow",_) $ (t1 as Free (str1,_)) $ 
neuper@37950
  1602
		      (t2 as Free (str2,_))) v :mv_poly option= 
neuper@37950
  1603
    let
neuper@38006
  1604
	val x= Unsynchronized.ref  NONE;
neuper@38006
  1605
	val len= Unsynchronized.ref  0;
neuper@38006
  1606
	val vl= Unsynchronized.ref  [];
neuper@38006
  1607
	val vh= Unsynchronized.ref  [];
neuper@38006
  1608
	val vtemp= Unsynchronized.ref  [];
neuper@38006
  1609
	val i= Unsynchronized.ref  0;	 
neuper@37950
  1610
    in
neuper@37950
  1611
    (
neuper@37950
  1612
     if (not o is_numeral) str1 andalso is_numeral str2 then
neuper@37950
  1613
	 (
neuper@37950
  1614
	  len:=length(v);
neuper@37950
  1615
	  vh:=v;
neuper@37950
  1616
neuper@37950
  1617
	  while ((!len)>(!i)) do
neuper@37950
  1618
	      (
neuper@37950
  1619
	       if str1=hd((!vh)) then
neuper@37950
  1620
		   (
neuper@37950
  1621
		    vl:=((the o int_of_str) str2)::(!vl)
neuper@37950
  1622
		    )
neuper@37950
  1623
	       else 
neuper@37950
  1624
		   (
neuper@37950
  1625
		    vl:=0::(!vl)
neuper@37950
  1626
		    );
neuper@37950
  1627
		   vh:=tl(!vh);
neuper@37950
  1628
		   i:=(!i)+1     
neuper@37950
  1629
		   );
neuper@37950
  1630
	      SOME [(1,rev(!vl))] handle _ => NONE
neuper@37950
  1631
	      )
neuper@38031
  1632
     else error ("RATIONALS_TERM2POLY_EXCEPTION 1: Invalid term")
neuper@37950
  1633
	 )
neuper@37950
  1634
    end
neuper@38014
  1635
  | term2poly' (Const ("Groups.plus_class.plus",_) $ t1 $ t2) v :mv_poly option = 
neuper@37950
  1636
    (
neuper@37950
  1637
     SOME ((the(term2poly' t1 v)) @ (the(term2poly' t2 v))) handle _ => NONE
neuper@37950
  1638
	 )
neuper@38014
  1639
  | term2poly' (Const ("Groups.minus_class.minus",_) $ t1 $ t2) v :mv_poly option = 
neuper@37950
  1640
    (
neuper@37950
  1641
     SOME ((the(term2poly' t1 v)) @ mv_skalar_mul((the(term2poly' t2 v)),~1)) handle _ => NONE
neuper@37950
  1642
	 )
neuper@38031
  1643
  | term2poly' (term) v = error ("RATIONALS_TERM2POLY_EXCEPTION 2: Invalid term");
neuper@37950
  1644
neuper@37950
  1645
(*. translates an Isabelle term into internal representation.
neuper@37950
  1646
    term2poly
neuper@37950
  1647
    fn : term ->              (*normalform [2]                    *)
neuper@37950
  1648
    	 string list ->       (*for ...!!! BITTE DIE ERKLÄRUNG, 
neuper@37950
  1649
    			       DIE DU MIR LETZTES MAL GEGEBEN HAST*)
neuper@37950
  1650
    	 mv_monom list        (*internal representation           *)
neuper@37950
  1651
    		  option      (*the translation may fail with NONE*)
neuper@37950
  1652
.*)
neuper@37950
  1653
fun term2poly (t:term) v = 
neuper@37950
  1654
     if is_polynomial t then term2poly' t v
neuper@38031
  1655
     else error ("term2poly: invalid = "^(term2str t));
neuper@37950
  1656
neuper@37950
  1657
(*. same as term2poly with automatic detection of the variables .*)
neuper@37950
  1658
fun term2polyx t = term2poly t (((map free2str) o vars) t); 
neuper@37950
  1659
neuper@37950
  1660
(*. checks if the term is in expanded polynomial form and converts it into the internal representation .*)
neuper@37950
  1661
fun expanded2poly (t:term) v = 
neuper@37950
  1662
    (*if is_expanded t then*) term2poly' t v
neuper@38031
  1663
    (*else error ("RATIONALS_EXPANDED2POLY_EXCEPTION: Invalid Polynomial")*);
neuper@37950
  1664
neuper@37950
  1665
(*. same as expanded2poly with automatic detection of the variables .*)
neuper@37950
  1666
fun expanded2polyx t = expanded2poly t (((map free2str) o vars) t);
neuper@37950
  1667
neuper@37950
  1668
(*. converts a powerproduct into term representation .*)
neuper@37950
  1669
fun powerproduct2term(xs,v) =  
neuper@37950
  1670
    let
neuper@38006
  1671
	val xss= Unsynchronized.ref  xs;
neuper@38006
  1672
	val vv= Unsynchronized.ref  v;
neuper@37950
  1673
    in
neuper@37950
  1674
	(
neuper@37950
  1675
	 while hd(!xss)=0 do 
neuper@37950
  1676
	     (
neuper@37950
  1677
	      xss:=tl(!xss);
neuper@37950
  1678
	      vv:=tl(!vv)
neuper@37950
  1679
	      );
neuper@37950
  1680
	     
neuper@37950
  1681
	 if list_is_null(tl(!xss)) then 
neuper@37950
  1682
	     (
neuper@37950
  1683
	      if hd(!xss)=1 then Free(hd(!vv), HOLogic.realT)
neuper@37950
  1684
	      else
neuper@37950
  1685
		  (
neuper@37950
  1686
		   Const("Atools.pow",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1687
		   Free(hd(!vv), HOLogic.realT) $  Free(str_of_int (hd(!xss)),HOLogic.realT)
neuper@37950
  1688
		   )
neuper@37950
  1689
	      )
neuper@37950
  1690
	 else
neuper@37950
  1691
	     (
neuper@37950
  1692
	      if hd(!xss)=1 then 
neuper@37950
  1693
		  ( 
neuper@38034
  1694
		   Const("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1695
		   Free(hd(!vv), HOLogic.realT) $
neuper@37950
  1696
		   powerproduct2term(tl(!xss),tl(!vv))
neuper@37950
  1697
		   )
neuper@37950
  1698
	      else
neuper@37950
  1699
		  (
neuper@38034
  1700
		   Const("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1701
		   (
neuper@37950
  1702
		    Const("Atools.pow",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1703
		    Free(hd(!vv), HOLogic.realT) $  Free(str_of_int (hd(!xss)),HOLogic.realT)
neuper@37950
  1704
		    ) $
neuper@37950
  1705
		    powerproduct2term(tl(!xss),tl(!vv))
neuper@37950
  1706
		   )
neuper@37950
  1707
	      )
neuper@37950
  1708
	 )
neuper@37950
  1709
    end;
neuper@37950
  1710
neuper@37950
  1711
(*. converts a monom into term representation .*)
neuper@37950
  1712
(*fun monom2term ((c,e):mv_monom, v:string list) = 
neuper@37950
  1713
    if c=0 then Free(str_of_int 0,HOLogic.realT)  
neuper@37950
  1714
    else
neuper@37950
  1715
	(
neuper@37950
  1716
	 if list_is_null(e) then
neuper@37950
  1717
	     ( 
neuper@37950
  1718
	      Free(str_of_int c,HOLogic.realT)  
neuper@37950
  1719
	      )
neuper@37950
  1720
	 else
neuper@37950
  1721
	     (
neuper@37950
  1722
	      if c=1 then 
neuper@37950
  1723
		  (
neuper@37950
  1724
		   powerproduct2term(e,v)
neuper@37950
  1725
		   )
neuper@37950
  1726
	      else
neuper@37950
  1727
		  (
neuper@38034
  1728
		   Const("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $
neuper@37950
  1729
		   Free(str_of_int c,HOLogic.realT)  $
neuper@37950
  1730
		   powerproduct2term(e,v)
neuper@37950
  1731
		   )
neuper@37950
  1732
		  )
neuper@37950
  1733
	     );*)
neuper@37950
  1734
neuper@37950
  1735
neuper@37950
  1736
(*fun monom2term ((i, is):mv_monom, v) = 
neuper@37950
  1737
    if list_is_null is 
neuper@37950
  1738
    then 
neuper@37950
  1739
	if i >= 0 
neuper@37950
  1740
	then Free (str_of_int i, HOLogic.realT)
neuper@37950
  1741
	else Const ("uminus", HOLogic.realT --> HOLogic.realT) $
neuper@37950
  1742
		   Free ((str_of_int o abs) i, HOLogic.realT)
neuper@37950
  1743
    else
neuper@37950
  1744
	if i > 0 
neuper@38034
  1745
	then Const ("Groups.times_class.times", [HOLogic.realT,HOLogic.realT]---> HOLogic.realT) $
neuper@37950
  1746
		   (Free (str_of_int i, HOLogic.realT)) $
neuper@37950
  1747
		   powerproduct2term(is, v)
neuper@38034
  1748
	else Const ("Groups.times_class.times", [HOLogic.realT,HOLogic.realT]---> HOLogic.realT) $
neuper@37950
  1749
		   (Const ("uminus", HOLogic.realT --> HOLogic.realT) $
neuper@37950
  1750
		   Free ((str_of_int o abs) i, HOLogic.realT)) $
neuper@37950
  1751
		   powerproduct2term(is, vs);---------------------------*)
neuper@37950
  1752
fun monom2term ((i, is) : mv_monom, vs) = 
neuper@37950
  1753
    if list_is_null is 
neuper@37950
  1754
    then Free (str_of_int i, HOLogic.realT)
neuper@37950
  1755
    else if i = 1
neuper@37950
  1756
    then powerproduct2term (is, vs)
neuper@38034
  1757
    else Const ("Groups.times_class.times", [HOLogic.realT, HOLogic.realT] ---> HOLogic.realT) $
neuper@37950
  1758
	       (Free (str_of_int i, HOLogic.realT)) $
neuper@37950
  1759
	       powerproduct2term (is, vs);
neuper@37950
  1760
    
neuper@37950
  1761
(*. converts the internal polynomial representation into an Isabelle term.*)
neuper@37950
  1762
fun poly2term' ([] : mv_poly, vs) = Free(str_of_int 0, HOLogic.realT)  
neuper@37950
  1763
  | poly2term' ([(c, e) : mv_monom], vs) = monom2term ((c, e), vs)
neuper@37950
  1764
  | poly2term' ((c, e) :: ces, vs) =  
neuper@38014
  1765
    Const("Groups.plus_class.plus", [HOLogic.realT, HOLogic.realT] ---> HOLogic.realT) $
neuper@37950
  1766
         poly2term (ces, vs) $ monom2term ((c, e), vs)
neuper@37950
  1767
and poly2term (xs, vs) = poly2term' (rev (sort (mv_geq LEX_) (xs)), vs);
neuper@37950
  1768
neuper@37950
  1769
neuper@37950
  1770
(*. converts a monom into term representation .*)
neuper@37950
  1771
(*. ignores the sign of the coefficients => use only for exp-poly functions .*)
neuper@37950
  1772
fun monom2term2((c,e):mv_monom, v:string list) =  
neuper@37950
  1773
    if c=0 then Free(str_of_int 0,HOLogic.realT)  
neuper@37950
  1774
    else
neuper@37950
  1775
	(
neuper@37950
  1776
	 if list_is_null(e) then
neuper@37950
  1777
	     ( 
neuper@37950
  1778
	      Free(str_of_int (abs(c)),HOLogic.realT)  
neuper@37950
  1779
	      )
neuper@37950
  1780
	 else
neuper@37950
  1781
	     (
neuper@37950
  1782
	      if abs(c)=1 then 
neuper@37950
  1783
		  (
neuper@37950
  1784
		   powerproduct2term(e,v)
neuper@37950
  1785
		   )
neuper@37950
  1786
	      else
neuper@37950
  1787
		  (
neuper@38034
  1788
		   Const("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $
neuper@37950
  1789
		   Free(str_of_int (abs(c)),HOLogic.realT)  $
neuper@37950
  1790
		   powerproduct2term(e,v)
neuper@37950
  1791
		   )
neuper@37950
  1792
		  )
neuper@37950
  1793
	     );
neuper@37950
  1794
neuper@37950
  1795
(*. converts the expanded polynomial representation into the term representation .*)
neuper@37950
  1796
fun exp2term' ([]:mv_poly,vars) =  Free(str_of_int 0,HOLogic.realT)  
neuper@37950
  1797
  | exp2term' ([(c,e)],vars) =     monom2term((c,e),vars) 			     
neuper@37950
  1798
  | exp2term' ((c1,e1)::others,vars) =  
neuper@37950
  1799
    if c1<0 then 	
neuper@38014
  1800
	Const("Groups.minus_class.minus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $
neuper@37950
  1801
	exp2term'(others,vars) $
neuper@37950
  1802
	( 
neuper@37950
  1803
	 monom2term2((c1,e1),vars)
neuper@37950
  1804
	 ) 
neuper@37950
  1805
    else
neuper@38014
  1806
	Const("Groups.plus_class.plus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $
neuper@37950
  1807
	exp2term'(others,vars) $
neuper@37950
  1808
	( 
neuper@37950
  1809
	 monom2term2((c1,e1),vars)
neuper@37950
  1810
	 );
neuper@37950
  1811
	
neuper@37950
  1812
(*. sorts the powerproduct by lexicographic termorder and converts them into 
neuper@37950
  1813
    a term in polynomial representation .*)
neuper@37950
  1814
fun poly2expanded (xs,vars) = exp2term'(rev(sort (mv_geq LEX_) (xs)),vars);
neuper@37950
  1815
neuper@37950
  1816
(*. converts a polynomial into expanded form .*)
neuper@37950
  1817
fun polynomial2expanded t =  
neuper@37950
  1818
    (let 
neuper@37950
  1819
	val vars=(((map free2str) o vars) t);
neuper@37950
  1820
    in
neuper@37950
  1821
	SOME (poly2expanded (the (term2poly t vars), vars))
neuper@37950
  1822
    end) handle _ => NONE;
neuper@37950
  1823
neuper@37950
  1824
(*. converts a polynomial into polynomial form .*)
neuper@37950
  1825
fun expanded2polynomial t =  
neuper@37950
  1826
    (let 
neuper@37950
  1827
	val vars=(((map free2str) o vars) t);
neuper@37950
  1828
    in
neuper@37950
  1829
	SOME (poly2term (the (expanded2poly t vars), vars))
neuper@37950
  1830
    end) handle _ => NONE;
neuper@37950
  1831
neuper@37950
  1832
neuper@37950
  1833
(*. calculates the greatest common divisor of numerator and denominator and seperates it from each .*)
neuper@48789
  1834
fun step_cancel (t as Const ("Fields.inverse_class.divide",_) $ p1 $ p2) = 
neuper@37950
  1835
    let
neuper@38006
  1836
	val p1' = Unsynchronized.ref [];
neuper@38006
  1837
	val p2' = Unsynchronized.ref [];
neuper@38006
  1838
	val p3  = Unsynchronized.ref []
neuper@37950
  1839
	val vars = rev(get_vars(p1) union get_vars(p2));
neuper@37950
  1840
    in
neuper@37950
  1841
	(
neuper@37950
  1842
         p1':= sort (mv_geq LEX_) (the (term2poly p1 vars ));
neuper@37950
  1843
       	 p2':= sort (mv_geq LEX_) (the (term2poly p2 vars ));
neuper@37950
  1844
	 p3:= sort (mv_geq LEX_) (mv_gcd (!p1') (!p2'));
neuper@37950
  1845
	 if (!p3)=[(1,mv_null2(vars))] then 
neuper@37950
  1846
	     (
neuper@48789
  1847
	      Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ p1 $ p2
neuper@37950
  1848
	      )
neuper@37950
  1849
	 else
neuper@37950
  1850
	     (
neuper@37950
  1851
neuper@37950
  1852
	      p1':=sort (mv_geq LEX_) (#1(mv_division((!p1'),(!p3),LEX_)));
neuper@37950
  1853
	      p2':=sort (mv_geq LEX_) (#1(mv_division((!p2'),(!p3),LEX_)));
neuper@37950
  1854
	      
neuper@37950
  1855
	      if #1(hd(sort (mv_geq LEX_) (!p2'))) (*mv_lc2(!p2',LEX_)*)>0 then
neuper@37950
  1856
	      (
neuper@48789
  1857
	       Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  1858
	       $ 
neuper@37950
  1859
	       (
neuper@38034
  1860
		Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1861
		poly2term(!p1',vars) $ 
neuper@37950
  1862
		poly2term(!p3,vars) 
neuper@37950
  1863
		) 
neuper@37950
  1864
	       $ 
neuper@37950
  1865
	       (
neuper@38034
  1866
		Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1867
		poly2term(!p2',vars) $ 
neuper@37950
  1868
		poly2term(!p3,vars)
neuper@37950
  1869
		) 	
neuper@37950
  1870
	       )	
neuper@37950
  1871
	      else
neuper@37950
  1872
	      (
neuper@37950
  1873
	       p1':=mv_skalar_mul(!p1',~1);
neuper@37950
  1874
	       p2':=mv_skalar_mul(!p2',~1);
neuper@37950
  1875
	       p3:=mv_skalar_mul(!p3,~1);
neuper@37950
  1876
	       (
neuper@48789
  1877
		Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  1878
		$ 
neuper@37950
  1879
		(
neuper@38034
  1880
		 Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1881
		 poly2term(!p1',vars) $ 
neuper@37950
  1882
		 poly2term(!p3,vars) 
neuper@37950
  1883
		 ) 
neuper@37950
  1884
		$ 
neuper@37950
  1885
		(
neuper@38034
  1886
		 Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1887
		 poly2term(!p2',vars) $ 
neuper@37950
  1888
		 poly2term(!p3,vars)
neuper@37950
  1889
		 ) 	
neuper@37950
  1890
		)	
neuper@37950
  1891
	       )	  
neuper@37950
  1892
	      )
neuper@37950
  1893
	     )
neuper@37950
  1894
    end
neuper@38031
  1895
| step_cancel _ = error ("RATIONALS_STEP_CANCEL_EXCEPTION: Invalid fraction"); 
neuper@37950
  1896
neuper@37950
  1897
neuper@37950
  1898
(*. same as step_cancel, this time for expanded forms (input+output) .*)
neuper@48789
  1899
fun step_cancel_expanded (t as Const ("Fields.inverse_class.divide",_) $ p1 $ p2) = 
neuper@37950
  1900
    let
neuper@38006
  1901
	val p1' = Unsynchronized.ref [];
neuper@38006
  1902
	val p2' = Unsynchronized.ref [];
neuper@38006
  1903
	val p3  = Unsynchronized.ref []
neuper@37950
  1904
	val vars = rev(get_vars(p1) union get_vars(p2));
neuper@37950
  1905
    in
neuper@37950
  1906
	(
neuper@37950
  1907
         p1':= sort (mv_geq LEX_) (the (expanded2poly p1 vars ));
neuper@37950
  1908
       	 p2':= sort (mv_geq LEX_) (the (expanded2poly p2 vars ));
neuper@37950
  1909
	 p3:= sort (mv_geq LEX_) (mv_gcd (!p1') (!p2'));
neuper@37950
  1910
	 if (!p3)=[(1,mv_null2(vars))] then 
neuper@37950
  1911
	     (
neuper@48789
  1912
	      Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ p1 $ p2
neuper@37950
  1913
	      )
neuper@37950
  1914
	 else
neuper@37950
  1915
	     (
neuper@37950
  1916
neuper@37950
  1917
	      p1':=sort (mv_geq LEX_) (#1(mv_division((!p1'),(!p3),LEX_)));
neuper@37950
  1918
	      p2':=sort (mv_geq LEX_) (#1(mv_division((!p2'),(!p3),LEX_)));
neuper@37950
  1919
	      
neuper@37950
  1920
	      if #1(hd(sort (mv_geq LEX_) (!p2')))(* mv_lc2(!p2',LEX_)*)>0 then
neuper@37950
  1921
	      (
neuper@48789
  1922
	       Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  1923
	       $ 
neuper@37950
  1924
	       (
neuper@38034
  1925
		Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1926
		poly2expanded(!p1',vars) $ 
neuper@37950
  1927
		poly2expanded(!p3,vars) 
neuper@37950
  1928
		) 
neuper@37950
  1929
	       $ 
neuper@37950
  1930
	       (
neuper@38034
  1931
		Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1932
		poly2expanded(!p2',vars) $ 
neuper@37950
  1933
		poly2expanded(!p3,vars)
neuper@37950
  1934
		) 	
neuper@37950
  1935
	       )	
neuper@37950
  1936
	      else
neuper@37950
  1937
	      (
neuper@37950
  1938
	       p1':=mv_skalar_mul(!p1',~1);
neuper@37950
  1939
	       p2':=mv_skalar_mul(!p2',~1);
neuper@37950
  1940
	       p3:=mv_skalar_mul(!p3,~1);
neuper@37950
  1941
	       (
neuper@48789
  1942
		Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  1943
		$ 
neuper@37950
  1944
		(
neuper@38034
  1945
		 Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1946
		 poly2expanded(!p1',vars) $ 
neuper@37950
  1947
		 poly2expanded(!p3,vars) 
neuper@37950
  1948
		 ) 
neuper@37950
  1949
		$ 
neuper@37950
  1950
		(
neuper@38034
  1951
		 Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  1952
		 poly2expanded(!p2',vars) $ 
neuper@37950
  1953
		 poly2expanded(!p3,vars)
neuper@37950
  1954
		 ) 	
neuper@37950
  1955
		)	
neuper@37950
  1956
	       )	  
neuper@37950
  1957
	      )
neuper@37950
  1958
	     )
neuper@37950
  1959
    end
neuper@38031
  1960
| step_cancel_expanded _ = error ("RATIONALS_STEP_CANCEL_EXCEPTION: Invalid fraction"); 
neuper@37950
  1961
neuper@37950
  1962
(*. calculates the greatest common divisor of numerator and denominator and divides each through it .*)
neuper@48789
  1963
fun direct_cancel (t as Const ("Fields.inverse_class.divide",_) $ p1 $ p2) = 
neuper@37950
  1964
    let
neuper@38006
  1965
	val p1' = Unsynchronized.ref [];
neuper@38006
  1966
	val p2' = Unsynchronized.ref [];
neuper@38006
  1967
	val p3  = Unsynchronized.ref []
neuper@37950
  1968
	val vars = rev(get_vars(p1) union get_vars(p2));
neuper@37950
  1969
    in
neuper@37950
  1970
	(
neuper@37950
  1971
	 p1':=sort (mv_geq LEX_) (mv_shorten((the (term2poly p1 vars )),LEX_));
neuper@37950
  1972
	 p2':=sort (mv_geq LEX_) (mv_shorten((the (term2poly p2 vars )),LEX_));	 
neuper@37950
  1973
	 p3 :=sort (mv_geq LEX_) (mv_gcd (!p1') (!p2'));
neuper@37950
  1974
neuper@37950
  1975
	 if (!p3)=[(1,mv_null2(vars))] then 
neuper@37950
  1976
	     (
neuper@48789
  1977
	      (Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ p1 $ p2,[])
neuper@37950
  1978
	      )
neuper@37950
  1979
	 else
neuper@37950
  1980
	     (
neuper@37950
  1981
	      p1':=sort (mv_geq LEX_) (#1(mv_division((!p1'),(!p3),LEX_)));
neuper@37950
  1982
	      p2':=sort (mv_geq LEX_) (#1(mv_division((!p2'),(!p3),LEX_)));
neuper@37950
  1983
	      if #1(hd(sort (mv_geq LEX_) (!p2'))) (*mv_lc2(!p2',LEX_)*)>0 then	      
neuper@37950
  1984
	      (
neuper@37950
  1985
	       (
neuper@48789
  1986
		Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  1987
		$ 
neuper@37950
  1988
		(
neuper@37950
  1989
		 poly2term((!p1'),vars)
neuper@37950
  1990
		 ) 
neuper@37950
  1991
		$ 
neuper@37950
  1992
		( 
neuper@37950
  1993
		 poly2term((!p2'),vars)
neuper@37950
  1994
		 ) 	
neuper@37950
  1995
		)
neuper@37950
  1996
	       ,
neuper@37950
  1997
	       if mv_grad(!p3)>0 then 
neuper@37950
  1998
		   [
neuper@37950
  1999
		    (
neuper@41929
  2000
		     Const ("HOL.Not",[bool]--->bool) $
neuper@37950
  2001
		     (
neuper@41922
  2002
		      Const("HOL.eq",[HOLogic.realT,HOLogic.realT]--->bool) $
neuper@37950
  2003
		      poly2term((!p3),vars) $
neuper@37950
  2004
		      Free("0",HOLogic.realT)
neuper@37950
  2005
		      )
neuper@37950
  2006
		     )
neuper@37950
  2007
		    ]
neuper@37950
  2008
	       else
neuper@37950
  2009
		   []
neuper@37950
  2010
		   )
neuper@37950
  2011
	      else
neuper@37950
  2012
		  (
neuper@37950
  2013
		   p1':=mv_skalar_mul(!p1',~1);
neuper@37950
  2014
		   p2':=mv_skalar_mul(!p2',~1);
neuper@37950
  2015
		   if length(!p3)> 2*(count_neg(!p3)) then () else p3 :=mv_skalar_mul(!p3,~1); 
neuper@37950
  2016
		       (
neuper@48789
  2017
			Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2018
			$ 
neuper@37950
  2019
			(
neuper@37950
  2020
			 poly2term((!p1'),vars)
neuper@37950
  2021
			 ) 
neuper@37950
  2022
			$ 
neuper@37950
  2023
			( 
neuper@37950
  2024
			 poly2term((!p2'),vars)
neuper@37950
  2025
			 ) 	
neuper@37950
  2026
			,
neuper@37950
  2027
			if mv_grad(!p3)>0 then 
neuper@37950
  2028
			    [
neuper@37950
  2029
			     (
neuper@41929
  2030
			      Const ("HOL.Not",[bool]--->bool) $
neuper@37950
  2031
			      (
neuper@41922
  2032
			       Const("HOL.eq",[HOLogic.realT,HOLogic.realT]--->bool) $
neuper@37950
  2033
			       poly2term((!p3),vars) $
neuper@37950
  2034
			       Free("0",HOLogic.realT)
neuper@37950
  2035
			       )
neuper@37950
  2036
			      )
neuper@37950
  2037
			     ]
neuper@37950
  2038
			else
neuper@37950
  2039
			    []
neuper@37950
  2040
			    )
neuper@37950
  2041
		       )
neuper@37950
  2042
		  )
neuper@37950
  2043
	     )
neuper@37950
  2044
    end
neuper@38031
  2045
  | direct_cancel _ = error ("RATIONALS_DIRECT_CANCEL_EXCEPTION: Invalid fraction"); 
neuper@37950
  2046
neuper@37950
  2047
(*. same es direct_cancel, this time for expanded forms (input+output).*) 
neuper@48789
  2048
fun direct_cancel_expanded (t as Const ("Fields.inverse_class.divide",_) $ p1 $ p2) =  
neuper@37950
  2049
    let
neuper@38006
  2050
	val p1' = Unsynchronized.ref [];
neuper@38006
  2051
	val p2' = Unsynchronized.ref [];
neuper@38006
  2052
	val p3  = Unsynchronized.ref []
neuper@37950
  2053
	val vars = rev(get_vars(p1) union get_vars(p2));
neuper@37950
  2054
    in
neuper@37950
  2055
	(
neuper@37950
  2056
	 p1':=sort (mv_geq LEX_) (mv_shorten((the (expanded2poly p1 vars )),LEX_));
neuper@37950
  2057
	 p2':=sort (mv_geq LEX_) (mv_shorten((the (expanded2poly p2 vars )),LEX_));	 
neuper@37950
  2058
	 p3 :=sort (mv_geq LEX_) (mv_gcd (!p1') (!p2'));
neuper@37950
  2059
neuper@37950
  2060
	 if (!p3)=[(1,mv_null2(vars))] then 
neuper@37950
  2061
	     (
neuper@48789
  2062
	      (Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ p1 $ p2,[])
neuper@37950
  2063
	      )
neuper@37950
  2064
	 else
neuper@37950
  2065
	     (
neuper@37950
  2066
	      p1':=sort (mv_geq LEX_) (#1(mv_division((!p1'),(!p3),LEX_)));
neuper@37950
  2067
	      p2':=sort (mv_geq LEX_) (#1(mv_division((!p2'),(!p3),LEX_)));
neuper@37950
  2068
	      if #1(hd(sort (mv_geq LEX_) (!p2'))) (*mv_lc2(!p2',LEX_)*)>0 then	      
neuper@37950
  2069
	      (
neuper@37950
  2070
	       (
neuper@48789
  2071
		Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2072
		$ 
neuper@37950
  2073
		(
neuper@37950
  2074
		 poly2expanded((!p1'),vars)
neuper@37950
  2075
		 ) 
neuper@37950
  2076
		$ 
neuper@37950
  2077
		( 
neuper@37950
  2078
		 poly2expanded((!p2'),vars)
neuper@37950
  2079
		 ) 	
neuper@37950
  2080
		)
neuper@37950
  2081
	       ,
neuper@37950
  2082
	       if mv_grad(!p3)>0 then 
neuper@37950
  2083
		   [
neuper@37950
  2084
		    (
neuper@41929
  2085
		     Const ("HOL.Not",[bool]--->bool) $
neuper@37950
  2086
		     (
neuper@41922
  2087
		      Const("HOL.eq",[HOLogic.realT,HOLogic.realT]--->bool) $
neuper@37950
  2088
		      poly2expanded((!p3),vars) $
neuper@37950
  2089
		      Free("0",HOLogic.realT)
neuper@37950
  2090
		      )
neuper@37950
  2091
		     )
neuper@37950
  2092
		    ]
neuper@37950
  2093
	       else
neuper@37950
  2094
		   []
neuper@37950
  2095
		   )
neuper@37950
  2096
	      else
neuper@37950
  2097
		  (
neuper@37950
  2098
		   p1':=mv_skalar_mul(!p1',~1);
neuper@37950
  2099
		   p2':=mv_skalar_mul(!p2',~1);
neuper@37950
  2100
		   if length(!p3)> 2*(count_neg(!p3)) then () else p3 :=mv_skalar_mul(!p3,~1); 
neuper@37950
  2101
		       (
neuper@48789
  2102
			Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2103
			$ 
neuper@37950
  2104
			(
neuper@37950
  2105
			 poly2expanded((!p1'),vars)
neuper@37950
  2106
			 ) 
neuper@37950
  2107
			$ 
neuper@37950
  2108
			( 
neuper@37950
  2109
			 poly2expanded((!p2'),vars)
neuper@37950
  2110
			 ) 	
neuper@37950
  2111
			,
neuper@37950
  2112
			if mv_grad(!p3)>0 then 
neuper@37950
  2113
			    [
neuper@37950
  2114
			     (
neuper@41929
  2115
			      Const ("HOL.Not",[bool]--->bool) $
neuper@37950
  2116
			      (
neuper@41922
  2117
			       Const("HOL.eq",[HOLogic.realT,HOLogic.realT]--->bool) $
neuper@37950
  2118
			       poly2expanded((!p3),vars) $
neuper@37950
  2119
			       Free("0",HOLogic.realT)
neuper@37950
  2120
			       )
neuper@37950
  2121
			      )
neuper@37950
  2122
			     ]
neuper@37950
  2123
			else
neuper@37950
  2124
			    []
neuper@37950
  2125
			    )
neuper@37950
  2126
		       )
neuper@37950
  2127
		  )
neuper@37950
  2128
	     )
neuper@37950
  2129
    end
neuper@38031
  2130
  | direct_cancel_expanded _ = error ("RATIONALS_DIRECT_CANCEL_EXCEPTION: Invalid fraction"); 
neuper@37950
  2131
neuper@37950
  2132
neuper@37950
  2133
(*. adds two fractions .*)
neuper@48789
  2134
fun add_fract ((Const("Fields.inverse_class.divide",_) $ t11 $ t12),(Const("Fields.inverse_class.divide",_) $ t21 $ t22)) =
neuper@37950
  2135
    let
neuper@37950
  2136
	val vars=get_vars(t11) union get_vars(t12) union get_vars(t21) union get_vars(t22);
neuper@38006
  2137
	val t11'= Unsynchronized.ref  (the(term2poly t11 vars));
neuper@52070
  2138
(* stopped Test_Isac.thy ...
neuper@38015
  2139
val _= tracing"### add_fract: done t11"
neuper@52070
  2140
*)
neuper@38006
  2141
	val t12'= Unsynchronized.ref  (the(term2poly t12 vars));
neuper@52070
  2142
(* stopped Test_Isac.thy ...
neuper@38015
  2143
val _= tracing"### add_fract: done t12"
neuper@52070
  2144
*)
neuper@38006
  2145
	val t21'= Unsynchronized.ref  (the(term2poly t21 vars));
neuper@52070
  2146
(* stopped Test_Isac.thy ...
neuper@38015
  2147
val _= tracing"### add_fract: done t21"
neuper@52070
  2148
*)
neuper@38006
  2149
	val t22'= Unsynchronized.ref  (the(term2poly t22 vars));
neuper@52070
  2150
(* stopped Test_Isac.thy ...
neuper@38015
  2151
val _= tracing"### add_fract: done t22"
neuper@52070
  2152
*)
neuper@38006
  2153
	val den= Unsynchronized.ref  [];
neuper@38006
  2154
	val nom= Unsynchronized.ref  [];
neuper@38006
  2155
	val m1= Unsynchronized.ref  [];
neuper@38006
  2156
	val m2= Unsynchronized.ref  [];
neuper@37950
  2157
    in
neuper@37950
  2158
	
neuper@37950
  2159
	(
neuper@37950
  2160
	 den :=sort (mv_geq LEX_) (mv_lcm (!t12') (!t22'));
neuper@38015
  2161
tracing"### add_fract: done sort mv_lcm";
neuper@37950
  2162
	 m1  :=sort (mv_geq LEX_) (#1(mv_division(!den,!t12',LEX_)));
neuper@38015
  2163
tracing"### add_fract: done sort mv_division t12";
neuper@37950
  2164
	 m2  :=sort (mv_geq LEX_) (#1(mv_division(!den,!t22',LEX_)));
neuper@38015
  2165
tracing"### add_fract: done sort mv_division t22";
neuper@37950
  2166
	 nom :=sort (mv_geq LEX_) 
neuper@37950
  2167
		    (mv_shorten(mv_add(mv_mul(!t11',!m1,LEX_),
neuper@37950
  2168
				       mv_mul(!t21',!m2,LEX_),
neuper@37950
  2169
				       LEX_),
neuper@37950
  2170
				LEX_));
neuper@38015
  2171
tracing"### add_fract: done sort mv_add";
neuper@37950
  2172
	 (
neuper@48789
  2173
	  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2174
	  $ 
neuper@37950
  2175
	  (
neuper@37950
  2176
	   poly2term((!nom),vars)
neuper@37950
  2177
	   ) 
neuper@37950
  2178
	  $ 
neuper@37950
  2179
	  ( 
neuper@37950
  2180
	   poly2term((!den),vars)
neuper@37950
  2181
	   )	      
neuper@37950
  2182
	  )
neuper@37950
  2183
	 )	     
neuper@37950
  2184
    end 
neuper@38031
  2185
  | add_fract (_,_) = error ("RATIONALS_ADD_FRACTION_EXCEPTION: Invalid add_fraction call");
neuper@37950
  2186
neuper@37950
  2187
(*. adds two expanded fractions .*)
neuper@48789
  2188
fun add_fract_exp ((Const("Fields.inverse_class.divide",_) $ t11 $ t12),(Const("Fields.inverse_class.divide",_) $ t21 $ t22)) =
neuper@37950
  2189
    let
neuper@37950
  2190
	val vars=get_vars(t11) union get_vars(t12) union get_vars(t21) union get_vars(t22);
neuper@38006
  2191
	val t11'= Unsynchronized.ref  (the(expanded2poly t11 vars));
neuper@38006
  2192
	val t12'= Unsynchronized.ref  (the(expanded2poly t12 vars));
neuper@38006
  2193
	val t21'= Unsynchronized.ref  (the(expanded2poly t21 vars));
neuper@38006
  2194
	val t22'= Unsynchronized.ref  (the(expanded2poly t22 vars));
neuper@38006
  2195
	val den= Unsynchronized.ref  [];
neuper@38006
  2196
	val nom= Unsynchronized.ref  [];
neuper@38006
  2197
	val m1= Unsynchronized.ref  [];
neuper@38006
  2198
	val m2= Unsynchronized.ref  [];
neuper@37950
  2199
    in
neuper@37950
  2200
	
neuper@37950
  2201
	(
neuper@37950
  2202
	 den :=sort (mv_geq LEX_) (mv_lcm (!t12') (!t22'));
neuper@37950
  2203
	 m1  :=sort (mv_geq LEX_) (#1(mv_division(!den,!t12',LEX_)));
neuper@37950
  2204
	 m2  :=sort (mv_geq LEX_) (#1(mv_division(!den,!t22',LEX_)));
neuper@37950
  2205
	 nom :=sort (mv_geq LEX_) (mv_shorten(mv_add(mv_mul(!t11',!m1,LEX_),mv_mul(!t21',!m2,LEX_),LEX_),LEX_));
neuper@37950
  2206
	 (
neuper@48789
  2207
	  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2208
	  $ 
neuper@37950
  2209
	  (
neuper@37950
  2210
	   poly2expanded((!nom),vars)
neuper@37950
  2211
	   ) 
neuper@37950
  2212
	  $ 
neuper@37950
  2213
	  ( 
neuper@37950
  2214
	   poly2expanded((!den),vars)
neuper@37950
  2215
	   )	      
neuper@37950
  2216
	  )
neuper@37950
  2217
	 )	     
neuper@37950
  2218
    end 
neuper@38031
  2219
  | add_fract_exp (_,_) = error ("RATIONALS_ADD_FRACTION_EXP_EXCEPTION: Invalid add_fraction call");
neuper@37950
  2220
neuper@37950
  2221
(*. adds a list of terms .*)
neuper@37950
  2222
fun add_list_of_fractions []= (Free("0",HOLogic.realT),[])
neuper@37950
  2223
  | add_list_of_fractions [x]= direct_cancel x
neuper@37950
  2224
  | add_list_of_fractions (x::y::xs) = 
neuper@37950
  2225
    let
neuper@37950
  2226
	val (t1a,rest1)=direct_cancel(x);
neuper@38015
  2227
val _= tracing"### add_list_of_fractions xs: has done direct_cancel(x)";
neuper@37950
  2228
	val (t2a,rest2)=direct_cancel(y);
neuper@38015
  2229
val _= tracing"### add_list_of_fractions xs: has done direct_cancel(y)";
neuper@37950
  2230
	val (t3a,rest3)=(add_list_of_fractions (add_fract(t1a,t2a)::xs));
neuper@38015
  2231
val _= tracing"### add_list_of_fractions xs: has done add_list_of_fraction xs";
neuper@37950
  2232
	val (t4a,rest4)=direct_cancel(t3a);
neuper@38015
  2233
val _= tracing"### add_list_of_fractions xs: has done direct_cancel(t3a)";
neuper@37950
  2234
	val rest=rest1 union rest2 union rest3 union rest4;
neuper@37950
  2235
    in
neuper@38015
  2236
	(tracing"### add_list_of_fractions in";
neuper@37950
  2237
	 (
neuper@37950
  2238
	 (t4a,rest) 
neuper@37950
  2239
	 )
neuper@37950
  2240
	 )
neuper@37950
  2241
    end;
neuper@37950
  2242
neuper@37950
  2243
(*. adds a list of expanded terms .*)
neuper@37950
  2244
fun add_list_of_fractions_exp []= (Free("0",HOLogic.realT),[])
neuper@37950
  2245
  | add_list_of_fractions_exp [x]= direct_cancel_expanded x
neuper@37950
  2246
  | add_list_of_fractions_exp (x::y::xs) = 
neuper@37950
  2247
    let
neuper@37950
  2248
	val (t1a,rest1)=direct_cancel_expanded(x);
neuper@37950
  2249
	val (t2a,rest2)=direct_cancel_expanded(y);
neuper@37950
  2250
	val (t3a,rest3)=(add_list_of_fractions_exp (add_fract_exp(t1a,t2a)::xs));
neuper@37950
  2251
	val (t4a,rest4)=direct_cancel_expanded(t3a);
neuper@37950
  2252
	val rest=rest1 union rest2 union rest3 union rest4;
neuper@37950
  2253
    in
neuper@37950
  2254
	(
neuper@37950
  2255
	 (t4a,rest) 
neuper@37950
  2256
	 )
neuper@37950
  2257
    end;
neuper@37950
  2258
neuper@37950
  2259
(*. calculates the lcm of a list of mv_poly .*)
neuper@37950
  2260
fun calc_lcm ([x],var)= (x,var) 
neuper@37950
  2261
  | calc_lcm ((x::xs),var) = (mv_lcm x (#1(calc_lcm (xs,var))),var);
neuper@37950
  2262
neuper@37950
  2263
(*. converts a list of terms to a list of mv_poly .*)
neuper@37950
  2264
fun t2d([],_)=[] 
neuper@48789
  2265
  | t2d((t as (Const("Fields.inverse_class.divide",_) $ p1 $ p2))::xs,vars)= (the(term2poly p2 vars)) :: t2d(xs,vars); 
neuper@37950
  2266
neuper@37950
  2267
(*. same as t2d, this time for expanded forms .*)
neuper@37950
  2268
fun t2d_exp([],_)=[]  
neuper@48789
  2269
  | t2d_exp((t as (Const("Fields.inverse_class.divide",_) $ p1 $ p2))::xs,vars)= (the(expanded2poly p2 vars)) :: t2d_exp(xs,vars);
neuper@37950
  2270
neuper@37950
  2271
(*. converts a list of fract terms to a list of their denominators .*)
neuper@37950
  2272
fun termlist2denominators [] = ([],[])
neuper@37950
  2273
  | termlist2denominators xs = 
neuper@37950
  2274
    let	
neuper@38006
  2275
	val xxs= Unsynchronized.ref  xs;
neuper@38006
  2276
	val var= Unsynchronized.ref  [];
neuper@37950
  2277
    in
neuper@37950
  2278
	var:=[];
neuper@37950
  2279
	while length(!xxs)>0 do
neuper@37950
  2280
	    (
neuper@37950
  2281
	     let 
neuper@48789
  2282
		 val (t as Const ("Fields.inverse_class.divide",_) $ p1x $ p2x)=hd(!xxs);
neuper@37950
  2283
	     in
neuper@37950
  2284
		 (
neuper@37950
  2285
		  xxs:=tl(!xxs);
neuper@37950
  2286
		  var:=((get_vars(p2x)) union (get_vars(p1x)) union (!var))
neuper@37950
  2287
		  )
neuper@37950
  2288
	     end
neuper@37950
  2289
	     );
neuper@37950
  2290
	    (t2d(xs,!var),!var)
neuper@37950
  2291
    end;
neuper@37950
  2292
neuper@37950
  2293
(*. calculates the lcm of a list of mv_poly .*)
neuper@37950
  2294
fun calc_lcm ([x],var)= (x,var) 
neuper@37950
  2295
  | calc_lcm ((x::xs),var) = (mv_lcm x (#1(calc_lcm (xs,var))),var);
neuper@37950
  2296
neuper@37950
  2297
(*. converts a list of terms to a list of mv_poly .*)
neuper@37950
  2298
fun t2d([],_)=[] 
neuper@48789
  2299
  | t2d((t as (Const("Fields.inverse_class.divide",_) $ p1 $ p2))::xs,vars)= (the(term2poly p2 vars)) :: t2d(xs,vars); 
neuper@37950
  2300
neuper@37950
  2301
(*. same as t2d, this time for expanded forms .*)
neuper@37950
  2302
fun t2d_exp([],_)=[]  
neuper@48789
  2303
  | t2d_exp((t as (Const("Fields.inverse_class.divide",_) $ p1 $ p2))::xs,vars)= (the(expanded2poly p2 vars)) :: t2d_exp(xs,vars);
neuper@37950
  2304
neuper@37950
  2305
(*. converts a list of fract terms to a list of their denominators .*)
neuper@37950
  2306
fun termlist2denominators [] = ([],[])
neuper@37950
  2307
  | termlist2denominators xs = 
neuper@37950
  2308
    let	
neuper@38006
  2309
	val xxs= Unsynchronized.ref  xs;
neuper@38006
  2310
	val var= Unsynchronized.ref  [];
neuper@37950
  2311
    in
neuper@37950
  2312
	var:=[];
neuper@37950
  2313
	while length(!xxs)>0 do
neuper@37950
  2314
	    (
neuper@37950
  2315
	     let 
neuper@48789
  2316
		 val (t as Const ("Fields.inverse_class.divide",_) $ p1x $ p2x)=hd(!xxs);
neuper@37950
  2317
	     in
neuper@37950
  2318
		 (
neuper@37950
  2319
		  xxs:=tl(!xxs);
neuper@37950
  2320
		  var:=((get_vars(p2x)) union (get_vars(p1x)) union (!var))
neuper@37950
  2321
		  )
neuper@37950
  2322
	     end
neuper@37950
  2323
	     );
neuper@37950
  2324
	    (t2d(xs,!var),!var)
neuper@37950
  2325
    end;
neuper@37950
  2326
neuper@37950
  2327
(*. same as termlist2denminators, this time for expanded forms .*)
neuper@37950
  2328
fun termlist2denominators_exp [] = ([],[])
neuper@37950
  2329
  | termlist2denominators_exp xs = 
neuper@37950
  2330
    let	
neuper@38006
  2331
	val xxs= Unsynchronized.ref  xs;
neuper@38006
  2332
	val var= Unsynchronized.ref  [];
neuper@37950
  2333
    in
neuper@37950
  2334
	var:=[];
neuper@37950
  2335
	while length(!xxs)>0 do
neuper@37950
  2336
	    (
neuper@37950
  2337
	     let 
neuper@48789
  2338
		 val (t as Const ("Fields.inverse_class.divide",_) $ p1x $ p2x)=hd(!xxs);
neuper@37950
  2339
	     in
neuper@37950
  2340
		 (
neuper@37950
  2341
		  xxs:=tl(!xxs);
neuper@37950
  2342
		  var:=((get_vars(p2x)) union (get_vars(p1x)) union (!var))
neuper@37950
  2343
		  )
neuper@37950
  2344
	     end
neuper@37950
  2345
	     );
neuper@37950
  2346
	    (t2d_exp(xs,!var),!var)
neuper@37950
  2347
    end;
neuper@37950
  2348
neuper@37950
  2349
(*. reduces all fractions to the least common denominator .*)
neuper@37950
  2350
fun com_den(x::xs,denom,den,var)=
neuper@37950
  2351
    let 
neuper@48789
  2352
	val (t as Const ("Fields.inverse_class.divide",_) $ p1' $ p2')=x;
neuper@37950
  2353
	val p2= sort (mv_geq LEX_) (the(term2poly p2' var));
neuper@37950
  2354
	val p3= #1(mv_division(denom,p2,LEX_));
neuper@37950
  2355
	val p1var=get_vars(p1');
neuper@37950
  2356
    in     
neuper@37950
  2357
	if length(xs)>0 then 
neuper@37950
  2358
	    if p3=[(1,mv_null2(var))] then
neuper@37950
  2359
		(
neuper@38014
  2360
		 Const ("Groups.plus_class.plus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT)
neuper@37950
  2361
		 $ 
neuper@37950
  2362
		 (
neuper@48789
  2363
		  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2364
		  $ 
neuper@37950
  2365
		  poly2term(the (term2poly p1' p1var),p1var)
neuper@37950
  2366
		  $ 
neuper@37950
  2367
		  den	
neuper@37950
  2368
		  )    
neuper@37950
  2369
		 $ 
neuper@37950
  2370
		 #1(com_den(xs,denom,den,var))
neuper@37950
  2371
		,
neuper@37950
  2372
		[]
neuper@37950
  2373
		)
neuper@37950
  2374
	    else
neuper@37950
  2375
		(
neuper@38014
  2376
		 Const ("Groups.plus_class.plus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2377
		 $ 
neuper@37950
  2378
		 (
neuper@48789
  2379
		  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2380
		  $ 
neuper@37950
  2381
		  (
neuper@38034
  2382
		   Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2383
		   poly2term(the (term2poly p1' p1var),p1var) $ 
neuper@37950
  2384
		   poly2term(p3,var)
neuper@37950
  2385
		   ) 
neuper@37950
  2386
		  $ 
neuper@37950
  2387
		  (
neuper@37950
  2388
		   den
neuper@37950
  2389
		   ) 	
neuper@37950
  2390
		  )
neuper@37950
  2391
		 $ 
neuper@37950
  2392
		 #1(com_den(xs,denom,den,var))
neuper@37950
  2393
		,
neuper@37950
  2394
		[]
neuper@37950
  2395
		)
neuper@37950
  2396
	else
neuper@37950
  2397
	    if p3=[(1,mv_null2(var))] then
neuper@37950
  2398
		(
neuper@37950
  2399
		 (
neuper@48789
  2400
		  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2401
		  $ 
neuper@37950
  2402
		  poly2term(the (term2poly p1' p1var),p1var)
neuper@37950
  2403
		  $ 
neuper@37950
  2404
		  den	
neuper@37950
  2405
		  )
neuper@37950
  2406
		 ,
neuper@37950
  2407
		 []
neuper@37950
  2408
		 )
neuper@37950
  2409
	     else
neuper@37950
  2410
		 (
neuper@48789
  2411
		  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2412
		  $ 
neuper@37950
  2413
		  (
neuper@38034
  2414
		   Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2415
		   poly2term(the (term2poly p1' p1var),p1var) $ 
neuper@37950
  2416
		   poly2term(p3,var)
neuper@37950
  2417
		   ) 
neuper@37950
  2418
		  $ 
neuper@37950
  2419
		  den 	
neuper@37950
  2420
		  ,
neuper@37950
  2421
		  []
neuper@37950
  2422
		  )
neuper@37950
  2423
    end;
neuper@37950
  2424
neuper@37950
  2425
(*. same as com_den, this time for expanded forms .*)
neuper@37950
  2426
fun com_den_exp(x::xs,denom,den,var)=
neuper@37950
  2427
    let 
neuper@48789
  2428
	val (t as Const ("Fields.inverse_class.divide",_) $ p1' $ p2')=x;
neuper@37950
  2429
	val p2= sort (mv_geq LEX_) (the(expanded2poly p2' var));
neuper@37950
  2430
	val p3= #1(mv_division(denom,p2,LEX_));
neuper@37950
  2431
	val p1var=get_vars(p1');
neuper@37950
  2432
    in     
neuper@37950
  2433
	if length(xs)>0 then 
neuper@37950
  2434
	    if p3=[(1,mv_null2(var))] then
neuper@37950
  2435
		(
neuper@38014
  2436
		 Const ("Groups.plus_class.plus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT)
neuper@37950
  2437
		 $ 
neuper@37950
  2438
		 (
neuper@48789
  2439
		  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2440
		  $ 
neuper@37950
  2441
		  poly2expanded(the(expanded2poly p1' p1var),p1var)
neuper@37950
  2442
		  $ 
neuper@37950
  2443
		  den	
neuper@37950
  2444
		  )    
neuper@37950
  2445
		 $ 
neuper@37950
  2446
		 #1(com_den_exp(xs,denom,den,var))
neuper@37950
  2447
		,
neuper@37950
  2448
		[]
neuper@37950
  2449
		)
neuper@37950
  2450
	    else
neuper@37950
  2451
		(
neuper@38014
  2452
		 Const ("Groups.plus_class.plus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2453
		 $ 
neuper@37950
  2454
		 (
neuper@48789
  2455
		  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2456
		  $ 
neuper@37950
  2457
		  (
neuper@38034
  2458
		   Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2459
		   poly2expanded(the(expanded2poly p1' p1var),p1var) $ 
neuper@37950
  2460
		   poly2expanded(p3,var)
neuper@37950
  2461
		   ) 
neuper@37950
  2462
		  $ 
neuper@37950
  2463
		  (
neuper@37950
  2464
		   den
neuper@37950
  2465
		   ) 	
neuper@37950
  2466
		  )
neuper@37950
  2467
		 $ 
neuper@37950
  2468
		 #1(com_den_exp(xs,denom,den,var))
neuper@37950
  2469
		,
neuper@37950
  2470
		[]
neuper@37950
  2471
		)
neuper@37950
  2472
	else
neuper@37950
  2473
	    if p3=[(1,mv_null2(var))] then
neuper@37950
  2474
		(
neuper@37950
  2475
		 (
neuper@48789
  2476
		  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2477
		  $ 
neuper@37950
  2478
		  poly2expanded(the(expanded2poly p1' p1var),p1var)
neuper@37950
  2479
		  $ 
neuper@37950
  2480
		  den	
neuper@37950
  2481
		  )
neuper@37950
  2482
		 ,
neuper@37950
  2483
		 []
neuper@37950
  2484
		 )
neuper@37950
  2485
	     else
neuper@37950
  2486
		 (
neuper@48789
  2487
		  Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
neuper@37950
  2488
		  $ 
neuper@37950
  2489
		  (
neuper@38034
  2490
		   Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2491
		   poly2expanded(the(expanded2poly p1' p1var),p1var) $ 
neuper@37950
  2492
		   poly2expanded(p3,var)
neuper@37950
  2493
		   ) 
neuper@37950
  2494
		  $ 
neuper@37950
  2495
		  den 	
neuper@37950
  2496
		  ,
neuper@37950
  2497
		  []
neuper@37950
  2498
		  )
neuper@37950
  2499
    end;
neuper@37950
  2500
neuper@37950
  2501
(* wird aktuell nicht mehr gebraucht, bei rückänderung schon 
neuper@37950
  2502
-------------------------------------------------------------
neuper@37950
  2503
(* WN0210???SK brauch ma des überhaupt *)
neuper@37950
  2504
fun com_den2(x::xs,denom,den,var)=
neuper@37950
  2505
    let 
neuper@48789
  2506
	val (t as Const ("Fields.inverse_class.divide",_) $ p1' $ p2')=x;
neuper@37950
  2507
	val p2= sort (mv_geq LEX_) (the(term2poly p2' var));
neuper@37950
  2508
	val p3= #1(mv_division(denom,p2,LEX_));
neuper@37950
  2509
	val p1var=get_vars(p1');
neuper@37950
  2510
    in     
neuper@37950
  2511
	if length(xs)>0 then 
neuper@37950
  2512
	    if p3=[(1,mv_null2(var))] then
neuper@37950
  2513
		(
neuper@38014
  2514
		 Const ("Groups.plus_class.plus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2515
		 poly2term(the(term2poly p1' p1var),p1var) $ 
neuper@37950
  2516
		 com_den2(xs,denom,den,var)
neuper@37950
  2517
		)
neuper@37950
  2518
	    else
neuper@37950
  2519
		(
neuper@38014
  2520
		 Const ("Groups.plus_class.plus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2521
		 (
neuper@37950
  2522
		   let 
neuper@37950
  2523
		       val p3'=poly2term(p3,var);
neuper@37950
  2524
		       val vars= (((map free2str) o vars) p1') union (((map free2str) o vars) p3');
neuper@37950
  2525
		   in
neuper@37950
  2526
		       poly2term(sort (mv_geq LEX_) (mv_mul(the(term2poly p1' vars) ,the(term2poly p3' vars),LEX_)),vars)
neuper@37950
  2527
		   end
neuper@37950
  2528
		  ) $ 
neuper@37950
  2529
		 com_den2(xs,denom,den,var)
neuper@37950
  2530
		)
neuper@37950
  2531
	else
neuper@37950
  2532
	    if p3=[(1,mv_null2(var))] then
neuper@37950
  2533
		(
neuper@37950
  2534
		 poly2term(the(term2poly p1' p1var),p1var)
neuper@37950
  2535
		 )
neuper@37950
  2536
	     else
neuper@37950
  2537
		 (
neuper@37950
  2538
		   let 
neuper@37950
  2539
		       val p3'=poly2term(p3,var);
neuper@37950
  2540
		       val vars= (((map free2str) o vars) p1') union (((map free2str) o vars) p3');
neuper@37950
  2541
		   in
neuper@37950
  2542
		       poly2term(sort (mv_geq LEX_) (mv_mul(the(term2poly p1' vars) ,the(term2poly p3' vars),LEX_)),vars)
neuper@37950
  2543
		   end
neuper@37950
  2544
		  )
neuper@37950
  2545
    end;
neuper@37950
  2546
neuper@37950
  2547
(* WN0210???SK brauch ma des überhaupt *)
neuper@37950
  2548
fun com_den_exp2(x::xs,denom,den,var)=
neuper@37950
  2549
    let 
neuper@48789
  2550
	val (t as Const ("Fields.inverse_class.divide",_) $ p1' $ p2')=x;
neuper@37950
  2551
	val p2= sort (mv_geq LEX_) (the(expanded2poly p2' var));
neuper@37950
  2552
	val p3= #1(mv_division(denom,p2,LEX_));
neuper@37950
  2553
	val p1var=get_vars p1';
neuper@37950
  2554
    in     
neuper@37950
  2555
	if length(xs)>0 then 
neuper@37950
  2556
	    if p3=[(1,mv_null2(var))] then
neuper@37950
  2557
		(
neuper@38014
  2558
		 Const ("Groups.plus_class.plus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2559
		 poly2expanded(the (expanded2poly p1' p1var),p1var) $ 
neuper@37950
  2560
		 com_den_exp2(xs,denom,den,var)
neuper@37950
  2561
		)
neuper@37950
  2562
	    else
neuper@37950
  2563
		(
neuper@38014
  2564
		 Const ("Groups.plus_class.plus",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2565
		 (
neuper@37950
  2566
		   let 
neuper@37950
  2567
		       val p3'=poly2expanded(p3,var);
neuper@37950
  2568
		       val vars= (((map free2str) o vars) p1') union (((map free2str) o vars) p3');
neuper@37950
  2569
		   in
neuper@37950
  2570
		       poly2expanded(sort (mv_geq LEX_) (mv_mul(the(expanded2poly p1' vars) ,the(expanded2poly p3' vars),LEX_)),vars)
neuper@37950
  2571
		   end
neuper@37950
  2572
		  ) $ 
neuper@37950
  2573
		 com_den_exp2(xs,denom,den,var)
neuper@37950
  2574
		)
neuper@37950
  2575
	else
neuper@37950
  2576
	    if p3=[(1,mv_null2(var))] then
neuper@37950
  2577
		(
neuper@37950
  2578
		 poly2expanded(the (expanded2poly p1' p1var),p1var)
neuper@37950
  2579
		 )
neuper@37950
  2580
	     else
neuper@37950
  2581
		 (
neuper@37950
  2582
		   let 
neuper@37950
  2583
		       val p3'=poly2expanded(p3,var);
neuper@37950
  2584
		       val vars= (((map free2str) o vars) p1') union (((map free2str) o vars) p3');
neuper@37950
  2585
		   in
neuper@37950
  2586
		       poly2expanded(sort (mv_geq LEX_) (mv_mul(the(expanded2poly p1' vars) ,the(expanded2poly p3' vars),LEX_)),vars)
neuper@37950
  2587
		   end
neuper@37950
  2588
		  )
neuper@37950
  2589
    end;
neuper@37950
  2590
---------------------------------------------------------*)
neuper@37950
  2591
neuper@37950
  2592
neuper@37950
  2593
(*. searches for an element y of a list ys, which has an gcd not 1 with x .*) 
neuper@37950
  2594
fun exists_gcd (x,[]) = false 
neuper@37950
  2595
  | exists_gcd (x,y::ys) = if mv_gcd x y = [(1,mv_null2(#2(hd(x))))] then  exists_gcd (x,ys)
neuper@37950
  2596
			   else true;
neuper@37950
  2597
neuper@37950
  2598
(*. divides each element of the list xs with y .*)
neuper@37950
  2599
fun list_div ([],y) = [] 
neuper@37950
  2600
  | list_div (x::xs,y) = 
neuper@37950
  2601
    let
neuper@37950
  2602
	val (d,r)=mv_division(x,y,LEX_);
neuper@37950
  2603
    in
neuper@37950
  2604
	if r=[] then 
neuper@37950
  2605
	    d::list_div(xs,y)
neuper@37950
  2606
	else x::list_div(xs,y)
neuper@37950
  2607
    end;
neuper@37950
  2608
    
neuper@37950
  2609
(*. checks if x is in the list ys .*)
neuper@37950
  2610
fun in_list (x,[]) = false 
neuper@37950
  2611
  | in_list (x,y::ys) = if x=y then true
neuper@37950
  2612
			else in_list(x,ys);
neuper@37950
  2613
neuper@37950
  2614
(*. deletes all equal elements of the list xs .*)
neuper@37950
  2615
fun kill_equal [] = [] 
neuper@37950
  2616
  | kill_equal (x::xs) = if in_list(x,xs) orelse x=[(1,mv_null2(#2(hd(x))))] then kill_equal(xs)
neuper@37950
  2617
			 else x::kill_equal(xs);
neuper@37950
  2618
neuper@37950
  2619
(*. searches for new factors .*)
neuper@37950
  2620
fun new_factors [] = []
neuper@37950
  2621
  | new_factors (list:mv_poly list):mv_poly list = 
neuper@37950
  2622
    let
neuper@37950
  2623
	val l = kill_equal list;
neuper@37950
  2624
	val len = length(l);
neuper@37950
  2625
    in
neuper@37950
  2626
	if len>=2 then
neuper@37950
  2627
	    (
neuper@37950
  2628
	     let
neuper@37950
  2629
		 val x::y::xs=l;
neuper@37950
  2630
		 val gcd=mv_gcd x y;
neuper@37950
  2631
	     in
neuper@37950
  2632
		 if gcd=[(1,mv_null2(#2(hd(x))))] then 
neuper@37950
  2633
		     ( 
neuper@37950
  2634
		      if exists_gcd(x,xs) then new_factors (y::xs @ [x])
neuper@37950
  2635
		      else x::new_factors(y::xs)
neuper@37950
  2636
	             )
neuper@37950
  2637
		 else gcd::new_factors(kill_equal(list_div(x::y::xs,gcd)))
neuper@37950
  2638
	     end
neuper@37950
  2639
	     )
neuper@37950
  2640
	else
neuper@37950
  2641
	    if len=1 then [hd(l)]
neuper@37950
  2642
	    else []
neuper@37950
  2643
    end;
neuper@37950
  2644
neuper@37950
  2645
(*. gets the factors of a list .*)
neuper@37950
  2646
fun get_factors x = new_factors x; 
neuper@37950
  2647
neuper@37950
  2648
(*. multiplies the elements of the list .*)
neuper@37950
  2649
fun multi_list [] = []
neuper@37950
  2650
  | multi_list (x::xs) = if xs=[] then x
neuper@37950
  2651
			 else mv_mul(x,multi_list xs,LEX_);
neuper@37950
  2652
neuper@37950
  2653
(*. makes a term out of the elements of the list (polynomial representation) .*)
neuper@37950
  2654
fun make_term ([],vars) = Free(str_of_int 0,HOLogic.realT) 
neuper@37950
  2655
  | make_term ((x::xs),vars) = if length(xs)=0 then poly2term(sort (mv_geq LEX_) (x),vars)
neuper@37950
  2656
			       else
neuper@37950
  2657
				   (
neuper@38034
  2658
				    Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2659
				    poly2term(sort (mv_geq LEX_) (x),vars) $ 
neuper@37950
  2660
				    make_term(xs,vars)
neuper@37950
  2661
				    );
neuper@37950
  2662
neuper@37950
  2663
(*. factorizes the denominator (polynomial representation) .*)				
neuper@37950
  2664
fun factorize_den (l,den,vars) = 
neuper@37950
  2665
    let
neuper@37950
  2666
	val factor_list=kill_equal( (get_factors l));
neuper@37950
  2667
	val mlist=multi_list(factor_list);
neuper@37950
  2668
	val (last,rest)=mv_division(den,multi_list(factor_list),LEX_);
neuper@37950
  2669
    in
neuper@37950
  2670
	if rest=[] then
neuper@37950
  2671
	    (
neuper@37950
  2672
	     if last=[(1,mv_null2(vars))] then make_term(factor_list,vars)
neuper@37950
  2673
	     else make_term(last::factor_list,vars)
neuper@37950
  2674
	     )
neuper@38031
  2675
	else error ("RATIONALS_FACTORIZE_DEN_EXCEPTION: Invalid factor by division")
neuper@37950
  2676
    end; 
neuper@37950
  2677
neuper@37950
  2678
(*. makes a term out of the elements of the list (expanded polynomial representation) .*)
neuper@37950
  2679
fun make_exp ([],vars) = Free(str_of_int 0,HOLogic.realT) 
neuper@37950
  2680
  | make_exp ((x::xs),vars) = if length(xs)=0 then poly2expanded(sort (mv_geq LEX_) (x),vars)
neuper@37950
  2681
			       else
neuper@37950
  2682
				   (
neuper@38034
  2683
				    Const ("Groups.times_class.times",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2684
				    poly2expanded(sort (mv_geq LEX_) (x),vars) $ 
neuper@37950
  2685
				    make_exp(xs,vars)
neuper@37950
  2686
				    );
neuper@37950
  2687
neuper@37950
  2688
(*. factorizes the denominator (expanded polynomial representation) .*)	
neuper@37950
  2689
fun factorize_den_exp (l,den,vars) = 
neuper@37950
  2690
    let
neuper@37950
  2691
	val factor_list=kill_equal( (get_factors l));
neuper@37950
  2692
	val mlist=multi_list(factor_list);
neuper@37950
  2693
	val (last,rest)=mv_division(den,multi_list(factor_list),LEX_);
neuper@37950
  2694
    in
neuper@37950
  2695
	if rest=[] then
neuper@37950
  2696
	    (
neuper@37950
  2697
	     if last=[(1,mv_null2(vars))] then make_exp(factor_list,vars)
neuper@37950
  2698
	     else make_exp(last::factor_list,vars)
neuper@37950
  2699
	     )
neuper@38031
  2700
	else error ("RATIONALS_FACTORIZE_DEN_EXP_EXCEPTION: Invalid factor by division")
neuper@37950
  2701
    end; 
neuper@37950
  2702
neuper@37950
  2703
(*. calculates the common denominator of all elements of the list and multiplies .*)
neuper@37950
  2704
(*. the nominators and denominators with the correct factor .*)
neuper@37950
  2705
(*. (polynomial representation) .*)
neuper@37950
  2706
fun step_add_list_of_fractions []=(Free("0",HOLogic.realT),[]:term list)
neuper@38031
  2707
  | step_add_list_of_fractions [x]= error ("RATIONALS_STEP_ADD_LIST_OF_FRACTIONS_EXCEPTION: Nothing to add")
neuper@37950
  2708
  | step_add_list_of_fractions (xs) = 
neuper@37950
  2709
    let
neuper@37950
  2710
        val den_list=termlist2denominators (xs); (* list of denominators *)
neuper@37950
  2711
	val (denom,var)=calc_lcm(den_list);      (* common denominator *)
neuper@37950
  2712
	val den=factorize_den(#1(den_list),denom,var); (* faktorisierter Nenner !!! *)
neuper@37950
  2713
    in
neuper@37950
  2714
	com_den(xs,denom,den,var)
neuper@37950
  2715
    end;
neuper@37950
  2716
neuper@37950
  2717
(*. calculates the common denominator of all elements of the list and multiplies .*)
neuper@37950
  2718
(*. the nominators and denominators with the correct factor .*)
neuper@37950
  2719
(*. (expanded polynomial representation) .*)
neuper@37950
  2720
fun step_add_list_of_fractions_exp []  = (Free("0",HOLogic.realT),[]:term list)
neuper@38031
  2721
  | step_add_list_of_fractions_exp [x] = error ("RATIONALS_STEP_ADD_LIST_OF_FRACTIONS_EXP_EXCEPTION: Nothing to add")
neuper@37950
  2722
  | step_add_list_of_fractions_exp (xs)= 
neuper@37950
  2723
    let
neuper@37950
  2724
        val den_list=termlist2denominators_exp (xs); (* list of denominators *)
neuper@37950
  2725
	val (denom,var)=calc_lcm(den_list);      (* common denominator *)
neuper@37950
  2726
	val den=factorize_den_exp(#1(den_list),denom,var); (* faktorisierter Nenner !!! *)
neuper@37950
  2727
    in
neuper@37950
  2728
	com_den_exp(xs,denom,den,var)
neuper@37950
  2729
    end;
neuper@37950
  2730
neuper@37950
  2731
(* wird aktuell nicht mehr gebraucht, bei rückänderung schon 
neuper@37950
  2732
-------------------------------------------------------------
neuper@37950
  2733
(* WN0210???SK brauch ma des überhaupt *)
neuper@37950
  2734
fun step_add_list_of_fractions2 []=(Free("0",HOLogic.realT),[]:term list)
neuper@37950
  2735
  | step_add_list_of_fractions2 [x]=(x,[])
neuper@37950
  2736
  | step_add_list_of_fractions2 (xs) = 
neuper@37950
  2737
    let
neuper@37950
  2738
        val den_list=termlist2denominators (xs); (* list of denominators *)
neuper@37950
  2739
	val (denom,var)=calc_lcm(den_list);      (* common denominator *)
neuper@37950
  2740
	val den=factorize_den(#1(den_list),denom,var);  (* faktorisierter Nenner !!! *)
neuper@37950
  2741
    in
neuper@37950
  2742
	(
neuper@48789
  2743
	 Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2744
	 com_den2(xs,denom, poly2term(denom,var)(*den*),var) $
neuper@37950
  2745
	 poly2term(denom,var)
neuper@37950
  2746
	,
neuper@37950
  2747
	[]
neuper@37950
  2748
	)
neuper@37950
  2749
    end;
neuper@37950
  2750
neuper@37950
  2751
(* WN0210???SK brauch ma des überhaupt *)
neuper@37950
  2752
fun step_add_list_of_fractions2_exp []=(Free("0",HOLogic.realT),[]:term list)
neuper@37950
  2753
  | step_add_list_of_fractions2_exp [x]=(x,[])
neuper@37950
  2754
  | step_add_list_of_fractions2_exp (xs) = 
neuper@37950
  2755
    let
neuper@37950
  2756
        val den_list=termlist2denominators_exp (xs); (* list of denominators *)
neuper@37950
  2757
	val (denom,var)=calc_lcm(den_list);      (* common denominator *)
neuper@37950
  2758
	val den=factorize_den_exp(#1(den_list),denom,var);  (* faktorisierter Nenner !!! *)
neuper@37950
  2759
    in
neuper@37950
  2760
	(
neuper@48789
  2761
	 Const ("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2762
	 com_den_exp2(xs,denom, poly2term(denom,var)(*den*),var) $
neuper@37950
  2763
	 poly2expanded(denom,var)
neuper@37950
  2764
	,
neuper@37950
  2765
	[]
neuper@37950
  2766
	)
neuper@37950
  2767
    end;
neuper@37950
  2768
---------------------------------------------- *)
neuper@37950
  2769
neuper@37950
  2770
neuper@41933
  2771
(* converts a term, which contains several terms seperated by +, into a list of these terms .*)
neuper@48789
  2772
fun term2list (t as (Const("Fields.inverse_class.divide",_) $ _ $ _)) = [t]
neuper@37950
  2773
  | term2list (t as (Const("Atools.pow",_) $ _ $ _)) = 
neuper@48789
  2774
      [Const ("Fields.inverse_class.divide", 
neuper@41933
  2775
        [HOLogic.realT,HOLogic.realT] ---> HOLogic.realT) $ 
neuper@37950
  2776
	  t $ Free("1",HOLogic.realT)
neuper@37950
  2777
     ]
neuper@37950
  2778
  | term2list (t as (Free(_,_))) = 
neuper@48789
  2779
    [Const("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2780
	  t $  Free("1",HOLogic.realT)
neuper@37950
  2781
     ]
neuper@38034
  2782
  | term2list (t as (Const("Groups.times_class.times",_) $ _ $ _)) = 
neuper@48789
  2783
    [Const("Fields.inverse_class.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
neuper@37950
  2784
	  t $ Free("1",HOLogic.realT)
neuper@37950
  2785
     ]
neuper@38014
  2786
  | term2list (Const("Groups.plus_class.plus",_) $ t1 $ t2) = term2list(t1) @ term2list(t2)
neuper@38014
  2787
  | term2list (Const("Groups.minus_class.minus",_) $ t1 $ t2) = 
neuper@38031
  2788
    error ("RATIONALS_TERM2LIST_EXCEPTION: - not implemented yet")
neuper@38031
  2789
  | term2list _ = error ("RATIONALS_TERM2LIST_EXCEPTION: invalid term");
neuper@37950
  2790
neuper@37950
  2791
(*.factors out the gcd of nominator and denominator:
neuper@37950
  2792
   a/b = (a' * gcd)/(b' * gcd),  a,b,gcd  are poly[2].*)
neuper@37950
  2793
fun factout_p_  (thy:theory) t = SOME (step_cancel t,[]:term list); 
neuper@37950
  2794
fun factout_ (thy:theory) t = SOME (step_cancel_expanded t,[]:term list); 
neuper@37950
  2795
neuper@37950
  2796
(*.cancels a single fraction with normalform [2]
neuper@37950
  2797
   resulting in a canceled fraction [2], see factout_ .*)
neuper@37950
  2798
fun cancel_p_ (thy:theory) t = (*WN.2.6.03 no rewrite -> NONE !*)
neuper@37950
  2799
    (let val (t',asm) = direct_cancel(*_expanded ... corrected MG.21.8.03*) t
neuper@37950
  2800
     in if t = t' then NONE else SOME (t',asm) 
neuper@37950
  2801
     end) handle _ => NONE;
neuper@37950
  2802
(*.the same as above with normalform [3]
neuper@37950
  2803
  val cancel_ :
neuper@37950
  2804
      theory ->        (*10.02 unused                                    *)
neuper@37950
  2805
      term -> 	       (*fraction in normalform [3]                      *)
neuper@37950
  2806
      (term * 	       (*fraction in normalform [3]                      *)
neuper@37950
  2807
       term list)      (*casual asumptions in normalform [3]             *)
neuper@37950
  2808
	  option       (*NONE: the function is not applicable            *).*)
neuper@37950
  2809
fun cancel_ (thy:theory) t = SOME (direct_cancel_expanded t) handle _ => NONE;
neuper@37950
  2810
neuper@37950
  2811
(*.transforms sums of at least 2 fractions [3] to
neuper@37950
  2812
   sums with the least common multiple as nominator.*)
neuper@37950
  2813
fun common_nominator_p_ (thy:theory) t =
neuper@38015
  2814
((*tracing("### common_nominator_p_ called");*)
neuper@37950
  2815
    SOME (step_add_list_of_fractions(term2list(t))) handle _ => NONE
neuper@37950
  2816
);
neuper@37950
  2817
fun common_nominator_ (thy:theory) t =
neuper@37950
  2818
    SOME (step_add_list_of_fractions_exp(term2list(t))) handle _ => NONE;
neuper@37950
  2819
neuper@37950
  2820
(*.add 2 or more fractions
neuper@37950
  2821
val add_fraction_p_ :
neuper@37950
  2822
      theory ->        (*10.02 unused                                    *)
neuper@37950
  2823
      term -> 	       (*2 or more fractions with normalform [2]         *)
neuper@37950
  2824
      (term * 	       (*one fraction with normalform [2]                *)
neuper@37950
  2825
       term list)      (*casual assumptions in normalform [2] WN0210???SK  *)
neuper@37950
  2826
	  option       (*NONE: the function is not applicable            *).*)
neuper@37950
  2827
fun add_fraction_p_ (thy:theory) t = 
neuper@38015
  2828
(tracing("### add_fraction_p_ called");
neuper@37950
  2829
    (let val ts = term2list t
neuper@37950
  2830
     in if 1 < length ts
neuper@37950
  2831
	then SOME (add_list_of_fractions ts)
neuper@38031
  2832
	else NONE (*error ("RATIONALS_ADD_EXCEPTION: nothing to add")*)
neuper@37950
  2833
     end) handle _ => NONE
neuper@37950
  2834
);
neuper@37950
  2835
(*.same as add_fraction_p_ but with normalform [3].*)
neuper@37950
  2836
(*SOME (step_add_list_of_fractions2(term2list(t))); *)
neuper@37950
  2837
fun add_fraction_ (thy:theory) t = 
neuper@41933
  2838
    if length (term2list t) > 1 
neuper@37950
  2839
    then SOME (add_list_of_fractions_exp(term2list(t))) handle _ => NONE
neuper@38031
  2840
    else (*error ("RATIONALS_ADD_FRACTION_EXCEPTION: nothing to add")*)
neuper@37950
  2841
	NONE;
neuper@37950
  2842
fun add_fraction_ (thy:theory) t = 
neuper@37950
  2843
    (if 1 < length (term2list t)
neuper@37950
  2844
     then SOME (add_list_of_fractions_exp (term2list t))
neuper@38031
  2845
     else (*error ("RATIONALS_ADD_FRACTION_EXCEPTION: nothing to add")*)
neuper@37950
  2846
	 NONE) handle _ => NONE;
neuper@37950
  2847
neuper@37950
  2848
(*SOME (step_add_list_of_fractions2_exp(term2list(t))); *)
neuper@37950
  2849
neuper@37950
  2850
(*. brings the term into a normal form .*)
neuper@37950
  2851
fun norm_rational_ (thy:theory) t = 
neuper@37950
  2852
    SOME (add_list_of_fractions(term2list(t))) handle _ => NONE; 
neuper@37950
  2853
fun norm_expanded_rat_ (thy:theory) t = 
neuper@37950
  2854
    SOME (add_list_of_fractions_exp(term2list(t))) handle _ => NONE; 
neuper@37950
  2855
neuper@37950
  2856
neuper@37950
  2857
(*.evaluates conditions in calculate_Rational.*)
neuper@37950
  2858
(*make local with FIXX@ME result:term *term list*)
neuper@37950
  2859
val calc_rat_erls = prep_rls(
neuper@37950
  2860
  Rls {id = "calc_rat_erls", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  2861
	 erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@37950
  2862
	 rules = 
neuper@41922
  2863
	 [Calc ("HOL.eq",eval_equal "#equal_"),
neuper@37950
  2864
	  Calc ("Atools.is'_const",eval_const "#is_const_"),
neuper@37978
  2865
	  Thm ("not_true",num_str @{thm not_true}),
neuper@37978
  2866
	  Thm ("not_false",num_str @{thm not_false})
neuper@37950
  2867
	  ], 
neuper@37950
  2868
	 scr = EmptyScr});
neuper@37950
  2869
neuper@37950
  2870
neuper@37950
  2871
(*.simplifies expressions with numerals;
neuper@37950
  2872
   does NOT rearrange the term by AC-rewriting; thus terms with variables 
neuper@37950
  2873
   need to have constants to be commuted together respectively.*)
neuper@42318
  2874
val calculate_Rational = prep_rls (merge_rls "calculate_Rational"
neuper@42318
  2875
	  (Rls {id = "divide", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  2876
	    erls = calc_rat_erls, srls = Erls,
neuper@42451
  2877
	    calc = [], errpatts = [],
neuper@42318
  2878
	    rules = 
neuper@48789
  2879
	      [Calc ("Fields.inverse_class.divide",eval_cancel "#divide_e"),
neuper@37950
  2880
	       
neuper@42318
  2881
	       Thm ("minus_divide_left",num_str (@{thm minus_divide_left} RS @{thm sym})),
neuper@42318
  2882
	         (*SYM - ?x / ?y = - (?x / ?y)  may come from subst*)
neuper@37950
  2883
	       
neuper@37969
  2884
	       Thm ("rat_add",num_str @{thm rat_add}),
neuper@42318
  2885
	         (*"[| a is_const; b is_const; c is_const; d is_const |] ==> \
neuper@42318
  2886
		           \a / c + b / d = (a * d) / (c * d) + (b * c ) / (d * c)"*)
neuper@37969
  2887
	       Thm ("rat_add1",num_str @{thm rat_add1}),
neuper@42318
  2888
	         (*"[| a is_const; b is_const; c is_const |] ==> a / c + b / c = (a + b) / c"*)
neuper@37969
  2889
	       Thm ("rat_add2",num_str @{thm rat_add2}),
neuper@42318
  2890
	         (*"[| ?a is_const; ?b is_const; ?c is_const |] ==> ?a / ?c + ?b = (?a + ?b * ?c) / ?c"*)
neuper@37969
  2891
	       Thm ("rat_add3",num_str @{thm rat_add3}),
neuper@42318
  2892
	         (*"[| a is_const; b is_const; c is_const |] ==> a + b / c = (a * c) / c + b / c"\
neuper@42318
  2893
		           .... is_const to be omitted here FIXME*)
neuper@37950
  2894
	       
neuper@42318
  2895
	       Thm ("rat_mult",num_str @{thm rat_mult}), 
neuper@42318
  2896
	         (*a / b * (c / d) = a * c / (b * d)*)
neuper@37965
  2897
	       Thm ("times_divide_eq_right",num_str @{thm times_divide_eq_right}),
neuper@42318
  2898
	         (*?x * (?y / ?z) = ?x * ?y / ?z*)
neuper@37965
  2899
	       Thm ("times_divide_eq_left",num_str @{thm times_divide_eq_left}),
neuper@42318
  2900
	         (*?y / ?z * ?x = ?y * ?x / ?z*)
neuper@37950
  2901
	       
neuper@37969
  2902
	       Thm ("real_divide_divide1",num_str @{thm real_divide_divide1}),
neuper@42318
  2903
	         (*"?y ~= 0 ==> ?u / ?v / (?y / ?z) = ?u / ?v * (?z / ?y)"*)
neuper@37965
  2904
	       Thm ("divide_divide_eq_left",num_str @{thm divide_divide_eq_left}),
neuper@42318
  2905
	         (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
neuper@37950
  2906
	       
neuper@37969
  2907
	       Thm ("rat_power", num_str @{thm rat_power}),
neuper@42318
  2908
	         (*"(?a / ?b) ^^^ ?n = ?a ^^^ ?n / ?b ^^^ ?n"*)
neuper@37950
  2909
	       
neuper@37969
  2910
	       Thm ("mult_cross",num_str @{thm mult_cross}),
neuper@42318
  2911
	         (*"[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)*)
neuper@37969
  2912
	       Thm ("mult_cross1",num_str @{thm mult_cross1}),
neuper@42318
  2913
	         (*"   b ~= 0            ==> (a / b = c    ) = (a     = b * c)*)
neuper@37969
  2914
	       Thm ("mult_cross2",num_str @{thm mult_cross2})
neuper@42318
  2915
	         (*"           d ~= 0    ==> (a     = c / d) = (a * d =     c)*)
neuper@37950
  2916
	       ], scr = EmptyScr})
neuper@42318
  2917
	  calculate_Poly);
neuper@37950
  2918
neuper@37950
  2919
(*("is_expanded", ("Rational.is'_expanded", eval_is_expanded ""))*)
neuper@37950
  2920
fun eval_is_expanded (thmid:string) _ 
neuper@37950
  2921
		       (t as (Const("Rational.is'_expanded", _) $ arg)) thy = 
neuper@37950
  2922
    if is_expanded arg
neuper@52070
  2923
    then SOME (mk_thmid thmid "" (term_to_string''' thy arg) "", 
neuper@52070
  2924
	         Trueprop $ (mk_equality (t, @{term True})))
neuper@52070
  2925
    else SOME (mk_thmid thmid "" (term_to_string''' thy arg) "", 
neuper@52070
  2926
	         Trueprop $ (mk_equality (t, @{term False})))
neuper@37950
  2927
  | eval_is_expanded _ _ _ _ = NONE; 
neuper@37950
  2928
neuper@37950
  2929
val rational_erls = 
neuper@37950
  2930
    merge_rls "rational_erls" calculate_Rational 
neuper@37950
  2931
	      (append_rls "is_expanded" Atools_erls 
neuper@37950
  2932
			  [Calc ("Rational.is'_expanded", eval_is_expanded "")
neuper@37950
  2933
			   ]);
neuper@37950
  2934
neuper@37950
  2935
neuper@37950
  2936
(*.3 'reverse-rewrite-sets' for symbolic computation on rationals:
neuper@37950
  2937
 =================================================================
neuper@37950
  2938
 A[2] 'cancel_p': .
neuper@37950
  2939
 A[3] 'cancel': .
neuper@37950
  2940
 B[2] 'common_nominator_p': transforms summands in a term [2]
neuper@37950
  2941
         to fractions with the (least) common multiple as nominator.
neuper@37950
  2942
 B[3] 'norm_rational': normalizes arbitrary algebraic terms (without 
neuper@37950
  2943
         radicals and transzendental functions) to one canceled fraction,
neuper@37950
  2944
	 nominator and denominator in polynomial form.
neuper@37950
  2945
neuper@37950
  2946
In order to meet isac's requirements for interactive and stepwise calculation,
neuper@37950
  2947
each 'reverse-rewerite-set' consists of an initialization for the interpreter 
neuper@37950
  2948
state and of 4 functions, each of which employs rewriting as much as possible.
neuper@37950
  2949
The signature of these functions are the same in each 'reverse-rewrite-set' 
neuper@37950
  2950
respectively.*)
neuper@37950
  2951
neuper@37950
  2952
(* ************************************************************************* *)
neuper@37950
  2953
neuper@37950
  2954
local(*. cancel_p
neuper@37950
  2955
------------------------
neuper@37950
  2956
cancels a single fraction consisting of two (uni- or multivariate)
neuper@37950
  2957
polynomials WN0609???SK[2] into another such a fraction; examples:
neuper@37950
  2958
neuper@37950
  2959
	   a^2 + -1*b^2         a + b
neuper@37950
  2960
        -------------------- = ---------
neuper@37950
  2961
	a^2 + -2*a*b + b^2     a + -1*b
neuper@37950
  2962
neuper@37950
  2963
        a^2    a
neuper@37950
  2964
        --- = ---
neuper@37950
  2965
         a     1
neuper@37950
  2966
neuper@37950
  2967
Remark: the reverse ruleset does _NOT_ work properly with other input !.*)
neuper@37950
  2968
(*WN020824 wir werden "uberlegen, wie wir ungeeignete inputs zur"uckweisen*)
neuper@37950
  2969
neuper@37950
  2970
val {rules, rew_ord=(_,ro),...} =
neuper@37950
  2971
    rep_rls (assoc_rls "make_polynomial");
neuper@37950
  2972
(*WN060829 ... make_deriv does not terminate with 1st expl above,
neuper@37950
  2973
           see rational.sml --- investigate rulesets for cancel_p ---*)
neuper@37950
  2974
val {rules, rew_ord=(_,ro),...} =
neuper@37950
  2975
    rep_rls (assoc_rls "rev_rew_p");
neuper@37950
  2976
neuper@37950
  2977
(*.init_state = fn : term -> istate
neuper@37950
  2978
initialzies the state of the script interpreter. The state is:
neuper@37950
  2979
neuper@37950
  2980
type rrlsstate =      (*state for reverse rewriting*)
neuper@37950
  2981
     (term *          (*the current formula*)
neuper@37950
  2982
      term *          (*the final term*)
neuper@37950
  2983
      rule list       (*'reverse rule list' (#)*)
neuper@37950
  2984
	    list *    (*may be serveral, eg. in norm_rational*)
neuper@37950
  2985
      (rule *         (*Thm (+ Thm generated from Calc) resulting in ...*)
neuper@37950
  2986
       (term *        (*... rewrite with ...*)
neuper@37950
  2987
	term list))   (*... assumptions*)
neuper@37950
  2988
	  list);      (*derivation from given term to normalform
neuper@37950
  2989
		       in reverse order with sym_thm;
neuper@37950
  2990
                       (#) could be extracted from here by (map #1)*).*)
neuper@37950
  2991
(* val {rules, rew_ord=(_,ro),...} =
neuper@37950
  2992
       rep_rls (assoc_rls "rev_rew_p")        (*USE ALWAYS, SEE val cancel_p*);
neuper@37972
  2993
   val (thy, eval_rls, ro) =(thy, Atools_erls, ro) (*..val cancel_p*);
neuper@37950
  2994
   val t = t;
neuper@37950
  2995
   *)
neuper@37950
  2996
fun init_state thy eval_rls ro t =
neuper@37950
  2997
    let val SOME (t',_) = factout_p_ thy t
neuper@37950
  2998
        val SOME (t'',asm) = cancel_p_ thy t
neuper@37950
  2999
        val der = reverse_deriv thy eval_rls rules ro NONE t'
neuper@37950
  3000
        val der = der @ [(Thm ("real_mult_div_cancel2",
neuper@37969
  3001
			       num_str @{thm real_mult_div_cancel2}),
neuper@37950
  3002
			  (t'',asm))]
neuper@37950
  3003
        val rs = (distinct_Thm o (map #1)) der
neuper@37950
  3004
	val rs = filter_out (eq_Thms ["sym_real_add_zero_left",
neuper@37950
  3005
				      "sym_real_mult_0",
neuper@37950
  3006
				      "sym_real_mult_1"
neuper@37950
  3007
				      (*..insufficient,eg.make_Polynomial*)])rs
neuper@37950
  3008
    in (t,t'',[rs(*here only _ONE_ to ease locate_rule*)],der) end;
neuper@37950
  3009
neuper@37950
  3010
(*.locate_rule = fn : rule list -> term -> rule
neuper@37950
  3011
		      -> (rule * (term * term list) option) list.
neuper@37950
  3012
  checks a rule R for being a cancel-rule, and if it is,
neuper@37950
  3013
  then return the list of rules (+ the terms they are rewriting to)
neuper@37950
  3014
  which need to be applied before R should be applied.
neuper@37950
  3015
  precondition: the rule is applicable to the argument-term.
neuper@37950
  3016
arguments:
neuper@37950
  3017
  rule list: the reverse rule list
neuper@37950
  3018
  -> term  : ... to which the rule shall be applied
neuper@37950
  3019
  -> rule  : ... to be applied to term
neuper@37950
  3020
value:
neuper@37950
  3021
  -> (rule           : a rule rewriting to ...
neuper@37950
  3022
      * (term        : ... the resulting term ...
neuper@37950
  3023
         * term list): ... with the assumptions ( //#0).
neuper@37950
  3024
      ) list         : there may be several such rules;
neuper@37950
  3025
		       the list is empty, if the rule has nothing to do
neuper@37950
  3026
		       with cancelation.*)
neuper@37950
  3027
(* val () = ();
neuper@37950
  3028
   *)
neuper@37950
  3029
fun locate_rule thy eval_rls ro [rs] t r =
neuper@37950
  3030
    if (id_of_thm r) mem (map (id_of_thm)) rs
neuper@37950
  3031
    then let val ropt =
neuper@37950
  3032
		 rewrite_ thy ro eval_rls true (thm_of_thm r) t;
neuper@37950
  3033
	 in case ropt of
neuper@37950
  3034
		SOME ta => [(r, ta)]
neuper@38015
  3035
	      | NONE => (tracing("### locate_rule:  rewrite "^
neuper@37950
  3036
				 (id_of_thm r)^" "^(term2str t)^" = NONE");
neuper@37950
  3037
			 []) end
neuper@38015
  3038
    else (tracing("### locate_rule:  "^(id_of_thm r)^" not mem rrls");[])
neuper@37950
  3039
  | locate_rule _ _ _ _ _ _ =
neuper@38031
  3040
    error ("locate_rule: doesnt match rev-sets in istate");
neuper@37950
  3041
neuper@37950
  3042
(*.next_rule = fn : rule list -> term -> rule option
neuper@37950
  3043
  for a given term return the next rules to be done for cancelling.
neuper@37950
  3044
arguments:
neuper@42451
  3045
  rule list     : the reverse rule list 
neuper@37950
  3046
  term          : the term for which ...
neuper@37950
  3047
value:
neuper@37950
  3048
  -> rule option: ... this rule is appropriate for cancellation;
neuper@37950
  3049
		  there may be no such rule (if the term is canceled already.*)
neuper@37972
  3050
(* val thy = thy;
neuper@37950
  3051
   val Rrls {rew_ord=(_,ro),...} = cancel;
neuper@37950
  3052
   val ([rs],t) = (rss,f);
neuper@37950
  3053
   next_rule thy eval_rls ro [rs] t;(*eval fun next_rule ... before!*)
neuper@37950
  3054
neuper@37972
  3055
   val (thy, [rs]) = (thy, revsets);
neuper@37950
  3056
   val Rrls {rew_ord=(_,ro),...} = cancel;
neuper@37950
  3057
   nex [rs] t;
neuper@37950
  3058
   *)
neuper@37950
  3059
fun next_rule thy eval_rls ro [rs] t =
neuper@37950
  3060
    let val der = make_deriv thy eval_rls rs ro NONE t;
neuper@37950
  3061
    in case der of
neuper@37950
  3062
(* val (_,r,_)::_ = der;
neuper@37950
  3063
   *)
neuper@37950
  3064
	   (_,r,_)::_ => SOME r
neuper@37950
  3065
	 | _ => NONE
neuper@37950
  3066
    end
neuper@37950
  3067
  | next_rule _ _ _ _ _ =
neuper@38031
  3068
    error ("next_rule: doesnt match rev-sets in istate");
neuper@37950
  3069
neuper@37950
  3070
(*.val attach_form = f : rule list -> term -> term
neuper@37950
  3071
			 -> (rule * (term * term list)) list
neuper@37950
  3072
  checks an input term TI, if it may belong to a current cancellation, by
neuper@37950
  3073
  trying to derive it from the given term TG.
neuper@37950
  3074
arguments:
neuper@37950
  3075
  term   : TG, the last one in the cancellation agreed upon by user + math-eng
neuper@37950
  3076
  -> term: TI, the next one input by the user
neuper@37950
  3077
value:
neuper@37950
  3078
  -> (rule           : the rule to be applied in order to reach TI
neuper@37950
  3079
      * (term        : ... obtained by applying the rule ...
neuper@37950
  3080
         * term list): ... and the respective assumptions.
neuper@37950
  3081
      ) list         : there may be several such rules;
neuper@37950
  3082
                       the list is empty, if the users term does not belong
neuper@37950
  3083
		       to a cancellation of the term last agreed upon.*)
neuper@37950
  3084
fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
neuper@37950
  3085
    []:(rule * (term * term list)) list;
neuper@37950
  3086
neuper@37950
  3087
in
neuper@37950
  3088
neuper@37950
  3089
val cancel_p =
neuper@37950
  3090
    Rrls {id = "cancel_p", prepat=[],
neuper@37950
  3091
	  rew_ord=("ord_make_polynomial",
neuper@37972
  3092
		   ord_make_polynomial false thy),
neuper@37950
  3093
	  erls = rational_erls,
neuper@38014
  3094
	  calc = [("PLUS"    ,("Groups.plus_class.plus"        ,eval_binop "#add_")),
neuper@38034
  3095
		  ("TIMES"   ,("Groups.times_class.times"        ,eval_binop "#mult_")),
neuper@48789
  3096
		  ("DIVIDE" ,("Fields.inverse_class.divide"  ,eval_cancel "#divide_e")),
neuper@37950
  3097
		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
neuper@42451
  3098
	  errpatts = [],
neuper@37950
  3099
	  scr=Rfuns {init_state  = init_state thy Atools_erls ro,
neuper@37950
  3100
		     normal_form = cancel_p_ thy,
neuper@37950
  3101
		     locate_rule = locate_rule thy Atools_erls ro,
neuper@37950
  3102
		     next_rule   = next_rule thy Atools_erls ro,
neuper@37950
  3103
		     attach_form = attach_form}}
neuper@37950
  3104
end;(*local*)
neuper@37950
  3105
neuper@37950
  3106
local(*.ad (1) 'cancel'
neuper@37950
  3107
------------------------------
neuper@37950
  3108
cancels a single fraction consisting of two (uni- or multivariate)
neuper@37950
  3109
polynomials WN0609???SK[3] into another such a fraction; examples:
neuper@37950
  3110
neuper@37950
  3111
	   a^2 - b^2           a + b
neuper@37950
  3112
        -------------------- = ---------
neuper@37950
  3113
	a^2 - 2*a*b + b^2      a - *b
neuper@37950
  3114
neuper@37950
  3115
Remark: the reverse ruleset does _NOT_ work properly with other input !.*)
neuper@37950
  3116
(*WN 24.8.02: wir werden "uberlegen, wie wir ungeeignete inputs zur"uckweisen*)
neuper@37950
  3117
neuper@37950
  3118
(*
neuper@37950
  3119
val SOME (Rls {rules=rules,rew_ord=(_,ro),...}) = 
neuper@37950
  3120
    assoc'(!ruleset',"expand_binoms");
neuper@37950
  3121
*)
neuper@37950
  3122
val {rules=rules,rew_ord=(_,ro),...} =
neuper@37950
  3123
    rep_rls (assoc_rls "expand_binoms");
neuper@37972
  3124
val thy = thy;
neuper@37950
  3125
neuper@37950
  3126
fun init_state thy eval_rls ro t =
neuper@37950
  3127
    let val SOME (t',_) = factout_ thy t;
neuper@37950
  3128
        val SOME (t'',asm) = cancel_ thy t;
neuper@37950
  3129
        val der = reverse_deriv thy eval_rls rules ro NONE t';
neuper@37950
  3130
        val der = der @ [(Thm ("real_mult_div_cancel2",
neuper@37969
  3131
			       num_str @{thm real_mult_div_cancel2}),
neuper@37950
  3132
			  (t'',asm))]
neuper@37950
  3133
        val rs = map #1 der;
neuper@37950
  3134
    in (t,t'',[rs],der) end;
neuper@37950
  3135
neuper@37950
  3136
fun locate_rule thy eval_rls ro [rs] t r =
neuper@37950
  3137
    if (id_of_thm r) mem (map (id_of_thm)) rs
neuper@37950
  3138
    then let val ropt = 
neuper@37950
  3139
		 rewrite_ thy ro eval_rls true (thm_of_thm r) t;
neuper@37950
  3140
	 in case ropt of
neuper@37950
  3141
		SOME ta => [(r, ta)]
neuper@38015
  3142
	      | NONE => (tracing("### locate_rule:  rewrite "^
neuper@37950
  3143
				 (id_of_thm r)^" "^(term2str t)^" = NONE");
neuper@37950
  3144
			 []) end
neuper@38015
  3145
    else (tracing("### locate_rule:  "^(id_of_thm r)^" not mem rrls");[])
neuper@37950
  3146
  | locate_rule _ _ _ _ _ _ = 
neuper@38031
  3147
    error ("locate_rule: doesnt match rev-sets in istate");
neuper@37950
  3148
neuper@37950
  3149
fun next_rule thy eval_rls ro [rs] t =
neuper@37950
  3150
    let val der = make_deriv thy eval_rls rs ro NONE t;
neuper@37950
  3151
    in case der of 
neuper@37950
  3152
(* val (_,r,_)::_ = der;
neuper@37950
  3153
   *)
neuper@37950
  3154
	   (_,r,_)::_ => SOME r
neuper@37950
  3155
	 | _ => NONE
neuper@37950
  3156
    end
neuper@37950
  3157
  | next_rule _ _ _ _ _ = 
neuper@38031
  3158
    error ("next_rule: doesnt match rev-sets in istate");
neuper@37950
  3159
neuper@37950
  3160
fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
neuper@37950
  3161
    []:(rule * (term * term list)) list;
neuper@37950
  3162
neuper@38036
  3163
(* pat matched with the current term gives an environment 
neuper@38036
  3164
   (or not: hen the Rrls not applied);
neuper@48760
  3165
   if pre1 and pre2 evaluate to @{term True} in this environment,
neuper@38036
  3166
   then the Rrls is applied. *)
neuper@38037
  3167
val pat = parse_patt thy "?r/?s :: real";
neuper@37979
  3168
val pre1 = parse_patt thy "?r is_expanded";
neuper@37979
  3169
val pre2 = parse_patt thy "?s is_expanded";
neuper@37950
  3170
val prepat = [([pre1, pre2], pat)];
neuper@37950
  3171
neuper@37950
  3172
in
neuper@37950
  3173
neuper@37950
  3174
neuper@37950
  3175
val cancel = 
neuper@37950
  3176
    Rrls {id = "cancel", prepat=prepat,
neuper@37950
  3177
	  rew_ord=("ord_make_polynomial",
neuper@37972
  3178
		   ord_make_polynomial false thy),
neuper@37950
  3179
	  erls = rational_erls, 
neuper@38014
  3180
	  calc = [("PLUS"    ,("Groups.plus_class.plus"        ,eval_binop "#add_")),
neuper@38034
  3181
		  ("TIMES"   ,("Groups.times_class.times"        ,eval_binop "#mult_")),
neuper@48789
  3182
		  ("DIVIDE" ,("Fields.inverse_class.divide"  ,eval_cancel "#divide_e")),
neuper@37950
  3183
		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
neuper@42451
  3184
          errpatts = [],
neuper@37950
  3185
	  scr=Rfuns {init_state  = init_state thy Atools_erls ro,
neuper@37950
  3186
		     normal_form = cancel_ thy, 
neuper@37950
  3187
		     locate_rule = locate_rule thy Atools_erls ro,
neuper@37950
  3188
		     next_rule   = next_rule thy Atools_erls ro,
neuper@37950
  3189
		     attach_form = attach_form}}
neuper@37950
  3190
end;(*local*)
neuper@37950
  3191
neuper@37950
  3192
local(*.ad [2] 'common_nominator_p'
neuper@37950
  3193
---------------------------------
neuper@37950
  3194
FIXME Beschreibung .*)
neuper@37950
  3195
neuper@37950
  3196
neuper@37950
  3197
val {rules=rules,rew_ord=(_,ro),...} =
neuper@37950
  3198
    rep_rls (assoc_rls "make_polynomial");
neuper@37950
  3199
(*WN060829 ... make_deriv does not terminate with 1st expl above,
neuper@37950
  3200
           see rational.sml --- investigate rulesets for cancel_p ---*)
neuper@37950
  3201
val {rules, rew_ord=(_,ro),...} =
neuper@37950
  3202
    rep_rls (assoc_rls "rev_rew_p");
neuper@37972
  3203
val thy = thy;
neuper@37950
  3204
neuper@37950
  3205
neuper@37950
  3206
(*.common_nominator_p_ = fn : theory -> term -> (term * term list) option
neuper@37950
  3207
  as defined above*)
neuper@37950
  3208
neuper@37950
  3209
(*.init_state = fn : term -> istate
neuper@37950
  3210
initialzies the state of the interactive interpreter. The state is:
neuper@37950
  3211
neuper@37950
  3212
type rrlsstate =      (*state for reverse rewriting*)
neuper@37950
  3213
     (term *          (*the current formula*)
neuper@37950
  3214
      term *          (*the final term*)
neuper@37950
  3215
      rule list       (*'reverse rule list' (#)*)
neuper@37950
  3216
	    list *    (*may be serveral, eg. in norm_rational*)
neuper@37950
  3217
      (rule *         (*Thm (+ Thm generated from Calc) resulting in ...*)
neuper@37950
  3218
       (term *        (*... rewrite with ...*)
neuper@37950
  3219
	term list))   (*... assumptions*)
neuper@37950
  3220
	  list);      (*derivation from given term to normalform
neuper@37950
  3221
		       in reverse order with sym_thm;
neuper@37950
  3222
                       (#) could be extracted from here by (map #1)*).*)
neuper@37950
  3223
fun init_state thy eval_rls ro t =
neuper@37950
  3224
    let val SOME (t',_) = common_nominator_p_ thy t;
neuper@37950
  3225
        val SOME (t'',asm) = add_fraction_p_ thy t;
neuper@37950
  3226
        val der = reverse_deriv thy eval_rls rules ro NONE t';
neuper@37950
  3227
        val der = der @ [(Thm ("real_mult_div_cancel2",
neuper@37969
  3228
			       num_str @{thm real_mult_div_cancel2}),
neuper@37950
  3229
			  (t'',asm))]
neuper@37950
  3230
        val rs = (distinct_Thm o (map #1)) der;
neuper@37950
  3231
	val rs = filter_out (eq_Thms ["sym_real_add_zero_left",
neuper@37950
  3232
				      "sym_real_mult_0",
neuper@37950
  3233
				      "sym_real_mult_1"]) rs;
neuper@37950
  3234
    in (t,t'',[rs(*here only _ONE_*)],der) end;
neuper@37950
  3235
neuper@37950
  3236
(* use"knowledge/Rational.ML";
neuper@37950
  3237
   *)
neuper@37950
  3238
neuper@37950
  3239
(*.locate_rule = fn : rule list -> term -> rule
neuper@37950
  3240
		      -> (rule * (term * term list) option) list.
neuper@37950
  3241
  checks a rule R for being a cancel-rule, and if it is,
neuper@37950
  3242
  then return the list of rules (+ the terms they are rewriting to)
neuper@37950
  3243
  which need to be applied before R should be applied.
neuper@37950
  3244
  precondition: the rule is applicable to the argument-term.
neuper@37950
  3245
arguments:
neuper@37950
  3246
  rule list: the reverse rule list
neuper@37950
  3247
  -> term  : ... to which the rule shall be applied
neuper@37950
  3248
  -> rule  : ... to be applied to term
neuper@37950
  3249
value:
neuper@37950
  3250
  -> (rule           : a rule rewriting to ...
neuper@37950
  3251
      * (term        : ... the resulting term ...
neuper@37950
  3252
         * term list): ... with the assumptions ( //#0).
neuper@37950
  3253
      ) list         : there may be several such rules;
neuper@37950
  3254
		       the list is empty, if the rule has nothing to do
neuper@37950
  3255
		       with cancelation.*)
neuper@37950
  3256
(* val () = ();
neuper@37950
  3257
   *)
neuper@37950
  3258
fun locate_rule thy eval_rls ro [rs] t r =
neuper@37950
  3259
    if (id_of_thm r) mem (map (id_of_thm)) rs
neuper@37950
  3260
    then let val ropt =
neuper@37950
  3261
		 rewrite_ thy ro eval_rls true (thm_of_thm r) t;
neuper@37950
  3262
	 in case ropt of
neuper@37950
  3263
		SOME ta => [(r, ta)]
neuper@38015
  3264
	      | NONE => (tracing("### locate_rule:  rewrite "^
neuper@37950
  3265
				 (id_of_thm r)^" "^(term2str t)^" = NONE");
neuper@37950
  3266
			 []) end
neuper@38015
  3267
    else (tracing("### locate_rule:  "^(id_of_thm r)^" not mem rrls");[])
neuper@37950
  3268
  | locate_rule _ _ _ _ _ _ =
neuper@38031
  3269
    error ("locate_rule: doesnt match rev-sets in istate");
neuper@37950
  3270
neuper@37950
  3271
(*.next_rule = fn : rule list -> term -> rule option
neuper@37950
  3272
  for a given term return the next rules to be done for cancelling.
neuper@37950
  3273
arguments:
neuper@37950
  3274
  rule list     : the reverse rule list
neuper@37950
  3275
  term          : the term for which ...
neuper@37950
  3276
value:
neuper@37950
  3277
  -> rule option: ... this rule is appropriate for cancellation;
neuper@37950
  3278
		  there may be no such rule (if the term is canceled already.*)
neuper@37972
  3279
(* val thy = thy;
neuper@37950
  3280
   val Rrls {rew_ord=(_,ro),...} = cancel;
neuper@37950
  3281
   val ([rs],t) = (rss,f);
neuper@37950
  3282
   next_rule thy eval_rls ro [rs] t;(*eval fun next_rule ... before!*)
neuper@37950
  3283
neuper@37972
  3284
   val (thy, [rs]) = (thy, revsets);
neuper@37950
  3285
   val Rrls {rew_ord=(_,ro),...} = cancel;
neuper@37950
  3286
   nex [rs] t;
neuper@37950
  3287
   *)
neuper@37950
  3288
fun next_rule thy eval_rls ro [rs] t =
neuper@37950
  3289
    let val der = make_deriv thy eval_rls rs ro NONE t;
neuper@37950
  3290
    in case der of
neuper@37950
  3291
(* val (_,r,_)::_ = der;
neuper@37950
  3292
   *)
neuper@37950
  3293
	   (_,r,_)::_ => SOME r
neuper@37950
  3294
	 | _ => NONE
neuper@37950
  3295
    end
neuper@37950
  3296
  | next_rule _ _ _ _ _ =
neuper@38031
  3297
    error ("next_rule: doesnt match rev-sets in istate");
neuper@37950
  3298
neuper@37950
  3299
(*.val attach_form = f : rule list -> term -> term
neuper@37950
  3300
			 -> (rule * (term * term list)) list
neuper@37950
  3301
  checks an input term TI, if it may belong to a current cancellation, by
neuper@37950
  3302
  trying to derive it from the given term TG.
neuper@37950
  3303
arguments:
neuper@37950
  3304
  term   : TG, the last one in the cancellation agreed upon by user + math-eng
neuper@37950
  3305
  -> term: TI, the next one input by the user
neuper@37950
  3306
value:
neuper@37950
  3307
  -> (rule           : the rule to be applied in order to reach TI
neuper@37950
  3308
      * (term        : ... obtained by applying the rule ...
neuper@37950
  3309
         * term list): ... and the respective assumptions.
neuper@37950
  3310
      ) list         : there may be several such rules;
neuper@37950
  3311
                       the list is empty, if the users term does not belong
neuper@37950
  3312
		       to a cancellation of the term last agreed upon.*)
neuper@37950
  3313
fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
neuper@37950
  3314
    []:(rule * (term * term list)) list;
neuper@37950
  3315
neuper@38036
  3316
(* if each pat* matches with the current term, the Rrls is applied
neuper@48760
  3317
   (there are no preconditions to be checked, they are @{term True}) *)
neuper@38037
  3318
val pat0 = parse_patt thy "?r/?s+?u/?v :: real";
neuper@38037
  3319
val pat1 = parse_patt thy "?r/?s+?u    :: real";
neuper@38037
  3320
val pat2 = parse_patt thy "?r   +?u/?v :: real";
neuper@48760
  3321
val prepat = [([@{term True}], pat0),
neuper@48760
  3322
	      ([@{term True}], pat1),
neuper@48760
  3323
	      ([@{term True}], pat2)];
neuper@37950
  3324
in
neuper@37950
  3325
neuper@37950
  3326
(*11.02 schnelle L"osung f"ur RL: Bruch auch gek"urzt;
neuper@37950
  3327
  besser w"are: auf 1 gemeinsamen Bruchstrich, Nenner und Z"ahler unvereinfacht
neuper@37950
  3328
  dh. wie common_nominator_p_, aber auf 1 Bruchstrich*)
neuper@37950
  3329
val common_nominator_p =
neuper@37950
  3330
    Rrls {id = "common_nominator_p", prepat=prepat,
neuper@37950
  3331
	  rew_ord=("ord_make_polynomial",
neuper@37972
  3332
		   ord_make_polynomial false thy),
neuper@37950
  3333
	  erls = rational_erls,
neuper@38014
  3334
	  calc = [("PLUS"    ,("Groups.plus_class.plus"        ,eval_binop "#add_")),
neuper@38034
  3335
		  ("TIMES"   ,("Groups.times_class.times"        ,eval_binop "#mult_")),
neuper@48789
  3336
		  ("DIVIDE" ,("Fields.inverse_class.divide"  ,eval_cancel "#divide_e")),
neuper@37950
  3337
		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
neuper@42451
  3338
	  errpatts = [],
neuper@37950
  3339
	  scr=Rfuns {init_state  = init_state thy Atools_erls ro,
neuper@37950
  3340
		     normal_form = add_fraction_p_ thy,(*FIXME.WN0211*)
neuper@37950
  3341
		     locate_rule = locate_rule thy Atools_erls ro,
neuper@37950
  3342
		     next_rule   = next_rule thy Atools_erls ro,
neuper@37950
  3343
		     attach_form = attach_form}}
neuper@37950
  3344
end;(*local*)
neuper@37950
  3345
neuper@37950
  3346
local(*.ad [2] 'common_nominator'
neuper@37950
  3347
---------------------------------
neuper@37950
  3348
FIXME Beschreibung .*)
neuper@37950
  3349
neuper@37950
  3350
neuper@37950
  3351
val {rules=rules,rew_ord=(_,ro),...} =
neuper@37950
  3352
    rep_rls (assoc_rls "make_polynomial");
neuper@37972
  3353
val thy = thy;
neuper@37950
  3354
neuper@37950
  3355
neuper@37950
  3356
(*.common_nominator_ = fn : theory -> term -> (term * term list) option
neuper@37950
  3357
  as defined above*)
neuper@37950
  3358
neuper@37950
  3359
(*.init_state = fn : term -> istate
neuper@37950
  3360
initialzies the state of the interactive interpreter. The state is:
neuper@37950
  3361
type rrlsstate =      (*state for reverse rewriting*)
neuper@37950
  3362
     (term *          (*the current formula*)
neuper@37950
  3363
      term *          (*the final term*)
neuper@37950
  3364
      rule list       (*'reverse rule list' (#)*)
neuper@37950
  3365
	    list *    (*may be serveral, eg. in norm_rational*)
neuper@37950
  3366
      (rule *         (*Thm (+ Thm generated from Calc) resulting in ...*)
neuper@37950
  3367
       (term *        (*... rewrite with ...*)
neuper@37950
  3368
	term list))   (*... assumptions*)
neuper@37950
  3369
	  list);      (*derivation from given term to normalform
neuper@37950
  3370
		       in reverse order with sym_thm;
neuper@37950
  3371
                       (#) could be extracted from here by (map #1)*).*)
neuper@37950
  3372
fun init_state thy eval_rls ro t =
neuper@37950
  3373
    let val SOME (t',_) = common_nominator_ thy t;
neuper@37950
  3374
        val SOME (t'',asm) = add_fraction_ thy t;
neuper@37950
  3375
        val der = reverse_deriv thy eval_rls rules ro NONE t';
neuper@37950
  3376
        val der = der @ [(Thm ("real_mult_div_cancel2",
neuper@37969
  3377
			       num_str @{thm real_mult_div_cancel2}),
neuper@37950
  3378
			  (t'',asm))]
neuper@37950
  3379
        val rs = (distinct_Thm o (map #1)) der;
neuper@37950
  3380
	val rs = filter_out (eq_Thms ["sym_real_add_zero_left",
neuper@37950
  3381
				      "sym_real_mult_0",
neuper@37950
  3382
				      "sym_real_mult_1"]) rs;
neuper@37950
  3383
    in (t,t'',[rs(*here only _ONE_*)],der) end;
neuper@37950
  3384
neuper@37950
  3385
(*.locate_rule = fn : rule list -> term -> rule
neuper@37950
  3386
		      -> (rule * (term * term list) option) list.
neuper@37950
  3387
  checks a rule R for being a cancel-rule, and if it is,
neuper@37950
  3388
  then return the list of rules (+ the terms they are rewriting to)
neuper@37950
  3389
  which need to be applied before R should be applied.
neuper@37950
  3390
  precondition: the rule is applicable to the argument-term.
neuper@37950
  3391
arguments:
neuper@37950
  3392
  rule list: the reverse rule list
neuper@37950
  3393
  -> term  : ... to which the rule shall be applied
neuper@37950
  3394
  -> rule  : ... to be applied to term
neuper@37950
  3395
value:
neuper@37950
  3396
  -> (rule           : a rule rewriting to ...
neuper@37950
  3397
      * (term        : ... the resulting term ...
neuper@37950
  3398
         * term list): ... with the assumptions ( //#0).
neuper@37950
  3399
      ) list         : there may be several such rules;
neuper@37950
  3400
		       the list is empty, if the rule has nothing to do
neuper@37950
  3401
		       with cancelation.*)
neuper@37950
  3402
(* val () = ();
neuper@37950
  3403
   *)
neuper@37950
  3404
fun locate_rule thy eval_rls ro [rs] t r =
neuper@37950
  3405
    if (id_of_thm r) mem (map (id_of_thm)) rs
neuper@37950
  3406
    then let val ropt =
neuper@37950
  3407
		 rewrite_ thy ro eval_rls true (thm_of_thm r) t;
neuper@37950
  3408
	 in case ropt of
neuper@37950
  3409
		SOME ta => [(r, ta)]
neuper@38015
  3410
	      | NONE => (tracing("### locate_rule:  rewrite "^
neuper@37950
  3411
				 (id_of_thm r)^" "^(term2str t)^" = NONE");
neuper@37950
  3412
			 []) end
neuper@38015
  3413
    else (tracing("### locate_rule:  "^(id_of_thm r)^" not mem rrls");[])
neuper@37950
  3414
  | locate_rule _ _ _ _ _ _ =
neuper@38031
  3415
    error ("locate_rule: doesnt match rev-sets in istate");
neuper@37950
  3416
neuper@37950
  3417
(*.next_rule = fn : rule list -> term -> rule option
neuper@37950
  3418
  for a given term return the next rules to be done for cancelling.
neuper@37950
  3419
arguments:
neuper@37950
  3420
  rule list     : the reverse rule list
neuper@37950
  3421
  term          : the term for which ...
neuper@37950
  3422
value:
neuper@37950
  3423
  -> rule option: ... this rule is appropriate for cancellation;
neuper@37950
  3424
		  there may be no such rule (if the term is canceled already.*)
neuper@37972
  3425
(* val thy = thy;
neuper@37950
  3426
   val Rrls {rew_ord=(_,ro),...} = cancel;
neuper@37950
  3427
   val ([rs],t) = (rss,f);
neuper@37950
  3428
   next_rule thy eval_rls ro [rs] t;(*eval fun next_rule ... before!*)
neuper@37950
  3429
neuper@37972
  3430
   val (thy, [rs]) = (thy, revsets);
neuper@37950
  3431
   val Rrls {rew_ord=(_,ro),...} = cancel_p;
neuper@37950
  3432
   nex [rs] t;
neuper@37950
  3433
   *)
neuper@37950
  3434
fun next_rule thy eval_rls ro [rs] t =
neuper@37950
  3435
    let val der = make_deriv thy eval_rls rs ro NONE t;
neuper@37950
  3436
    in case der of
neuper@37950
  3437
(* val (_,r,_)::_ = der;
neuper@37950
  3438
   *)
neuper@37950
  3439
	   (_,r,_)::_ => SOME r
neuper@37950
  3440
	 | _ => NONE
neuper@37950
  3441
    end
neuper@37950
  3442
  | next_rule _ _ _ _ _ =
neuper@38031
  3443
    error ("next_rule: doesnt match rev-sets in istate");
neuper@37950
  3444
neuper@37950
  3445
(*.val attach_form = f : rule list -> term -> term
neuper@37950
  3446
			 -> (rule * (term * term list)) list
neuper@37950
  3447
  checks an input term TI, if it may belong to a current cancellation, by
neuper@37950
  3448
  trying to derive it from the given term TG.
neuper@37950
  3449
arguments:
neuper@37950
  3450
  term   : TG, the last one in the cancellation agreed upon by user + math-eng
neuper@37950
  3451
  -> term: TI, the next one input by the user
neuper@37950
  3452
value:
neuper@37950
  3453
  -> (rule           : the rule to be applied in order to reach TI
neuper@37950
  3454
      * (term        : ... obtained by applying the rule ...
neuper@37950
  3455
         * term list): ... and the respective assumptions.
neuper@37950
  3456
      ) list         : there may be several such rules;
neuper@37950
  3457
                       the list is empty, if the users term does not belong
neuper@37950
  3458
		       to a cancellation of the term last agreed upon.*)
neuper@37950
  3459
fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
neuper@37950
  3460
    []:(rule * (term * term list)) list;
neuper@37950
  3461
neuper@38036
  3462
(* if each pat* matches with the current term, the Rrls is applied
neuper@48760
  3463
   (there are no preconditions to be checked, they are @{term True}) *)
neuper@38037
  3464
val pat0 =  parse_patt thy "?r/?s+?u/?v :: real";
neuper@38037
  3465
val pat01 = parse_patt thy "?r/?s-?u/?v :: real";
neuper@38037
  3466
val pat1 =  parse_patt thy "?r/?s+?u    :: real";
neuper@38037
  3467
val pat11 = parse_patt thy "?r/?s-?u    :: real";
neuper@38037
  3468
val pat2 =  parse_patt thy "?r   +?u/?v :: real";
neuper@38037
  3469
val pat21 = parse_patt thy "?r   -?u/?v :: real";
neuper@48760
  3470
val prepat = [([@{term True}], pat0),
neuper@48760
  3471
	      ([@{term True}], pat01),
neuper@48760
  3472
	      ([@{term True}], pat1),
neuper@48760
  3473
	      ([@{term True}], pat11),
neuper@48760
  3474
	      ([@{term True}], pat2),
neuper@48760
  3475
	      ([@{term True}], pat21)];
neuper@37950
  3476
neuper@37950
  3477
neuper@37950
  3478
in
neuper@37950
  3479
neuper@37950
  3480
val common_nominator =
neuper@37950
  3481
    Rrls {id = "common_nominator", prepat=prepat,
neuper@37950
  3482
	  rew_ord=("ord_make_polynomial",
neuper@37972
  3483
		   ord_make_polynomial false thy),
neuper@37950
  3484
	  erls = rational_erls,
neuper@38014
  3485
	  calc = [("PLUS"    ,("Groups.plus_class.plus"        ,eval_binop "#add_")),
neuper@38034
  3486
		  ("TIMES"   ,("Groups.times_class.times"        ,eval_binop "#mult_")),
neuper@48789
  3487
		  ("DIVIDE" ,("Fields.inverse_class.divide"  ,eval_cancel "#divide_e")),
neuper@37950
  3488
		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
neuper@42451
  3489
	  errpatts = [],
neuper@37950
  3490
	  scr=Rfuns {init_state  = init_state thy Atools_erls ro,
neuper@37950
  3491
		     normal_form = add_fraction_ (*NOT common_nominator_*) thy,
neuper@37950
  3492
		     locate_rule = locate_rule thy Atools_erls ro,
neuper@37950
  3493
		     next_rule   = next_rule thy Atools_erls ro,
neuper@37950
  3494
		     attach_form = attach_form}}
neuper@37950
  3495
neuper@42451
  3496
end; (*local*)
neuper@42451
  3497
end; (*struct*)
neuper@42451
  3498
neuper@42451
  3499
*}
neuper@42451
  3500
ML {* 
neuper@37950
  3501
open RationalI;
neuper@37950
  3502
(*##*)
neuper@37950
  3503
neuper@37950
  3504
(*.the expression contains + - * ^ / only ?.*)
neuper@37950
  3505
fun is_ratpolyexp (Free _) = true
neuper@38014
  3506
  | is_ratpolyexp (Const ("Groups.plus_class.plus",_) $ Free _ $ Free _) = true
neuper@38014
  3507
  | is_ratpolyexp (Const ("Groups.minus_class.minus",_) $ Free _ $ Free _) = true
neuper@38034
  3508
  | is_ratpolyexp (Const ("Groups.times_class.times",_) $ Free _ $ Free _) = true
neuper@37950
  3509
  | is_ratpolyexp (Const ("Atools.pow",_) $ Free _ $ Free _) = true
neuper@48789
  3510
  | is_ratpolyexp (Const ("Fields.inverse_class.divide",_) $ Free _ $ Free _) = true
neuper@38014
  3511
  | is_ratpolyexp (Const ("Groups.plus_class.plus",_) $ t1 $ t2) = 
neuper@37950
  3512
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
neuper@38014
  3513
  | is_ratpolyexp (Const ("Groups.minus_class.minus",_) $ t1 $ t2) = 
neuper@37950
  3514
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
neuper@38034
  3515
  | is_ratpolyexp (Const ("Groups.times_class.times",_) $ t1 $ t2) = 
neuper@37950
  3516
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
neuper@37950
  3517
  | is_ratpolyexp (Const ("Atools.pow",_) $ t1 $ t2) = 
neuper@37950
  3518
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
neuper@48789
  3519
  | is_ratpolyexp (Const ("Fields.inverse_class.divide",_) $ t1 $ t2) = 
neuper@37950
  3520
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
neuper@37950
  3521
  | is_ratpolyexp _ = false;
neuper@37950
  3522
neuper@37950
  3523
(*("is_ratpolyexp", ("Rational.is'_ratpolyexp", eval_is_ratpolyexp ""))*)
neuper@37950
  3524
fun eval_is_ratpolyexp (thmid:string) _ 
neuper@37950
  3525
		       (t as (Const("Rational.is'_ratpolyexp", _) $ arg)) thy =
neuper@37950
  3526
    if is_ratpolyexp arg
neuper@52070
  3527
    then SOME (mk_thmid thmid "" (term_to_string''' thy arg) "", 
neuper@52070
  3528
	         Trueprop $ (mk_equality (t, @{term True})))
neuper@52070
  3529
    else SOME (mk_thmid thmid "" (term_to_string''' thy arg) "", 
neuper@52070
  3530
	         Trueprop $ (mk_equality (t, @{term False})))
neuper@37950
  3531
  | eval_is_ratpolyexp _ _ _ _ = NONE; 
neuper@37950
  3532
neuper@42301
  3533
(*("get_denominator", ("Rational.get_denominator", eval_get_denominator ""))*)
jan@42300
  3534
fun eval_get_denominator (thmid:string) _ 
neuper@42301
  3535
		      (t as Const ("Rational.get_denominator", _) $
neuper@48789
  3536
              (Const ("Fields.inverse_class.divide", _) $ num $
jan@42300
  3537
                denom)) thy = 
neuper@52070
  3538
      SOME (mk_thmid thmid "" (term_to_string''' thy denom) "", 
neuper@52070
  3539
	            Trueprop $ (mk_equality (t, denom)))
jan@42300
  3540
  | eval_get_denominator _ _ _ _ = NONE; 
neuper@37950
  3541
jan@42338
  3542
(*("get_numerator", ("Rational.get_numerator", eval_get_numerator ""))*)
jan@42338
  3543
fun eval_get_numerator (thmid:string) _ 
neuper@52070
  3544
      (t as Const ("Rational.get_numerator", _) $
neuper@52070
  3545
          (Const ("Fields.inverse_class.divide", _) $num
neuper@52070
  3546
            $denom )) thy = 
neuper@52070
  3547
    SOME (mk_thmid thmid "" (term_to_string''' thy num) "", 
neuper@52070
  3548
	    Trueprop $ (mk_equality (t, num)))
jan@42338
  3549
  | eval_get_numerator _ _ _ _ = NONE; 
jan@42338
  3550
neuper@37950
  3551
(*-------------------18.3.03 --> struct <-----------vvv--*)
neuper@37950
  3552
val add_fractions_p = common_nominator_p; (*FIXXXME:eilig f"ur norm_Rational*)
neuper@37950
  3553
neuper@37980
  3554
(*WN100906 removed in favour of discard_minus = discard_minus_ formerly 
neuper@37980
  3555
   discard binary minus, shift unary minus into -1*; 
neuper@37950
  3556
   unary minus before numerals are put into the numeral by parsing;
neuper@37980
  3557
   contains absolute minimum of thms for context in norm_Rational
neuper@42407
  3558
*** val discard_minus = prep_rls(
neuper@37950
  3559
  Rls {id = "discard_minus", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3560
      erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@42407
  3561
      rules =
neuper@42407
  3562
        [Thm ("real_diff_minus", num_str @{thm real_diff_minus}),
neuper@42407
  3563
	           (*"a - b = a + -1 * b"*)
neuper@42407
  3564
	         Thm ("sym_real_mult_minus1", num_str (@{thm real_mult_minus1} RS @{thm sym}))
neuper@42407
  3565
	           (*- ?z = "-1 * ?z"*)],
neuper@37979
  3566
      scr = EmptyScr
neuper@37950
  3567
      }):rls;
neuper@37980
  3568
 *)
neuper@37980
  3569
neuper@37950
  3570
(*erls for calculate_Rational; make local with FIXX@ME result:term *term list*)
neuper@37950
  3571
val powers_erls = prep_rls(
neuper@37950
  3572
  Rls {id = "powers_erls", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3573
      erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@37950
  3574
      rules = [Calc ("Atools.is'_atom",eval_is_atom "#is_atom_"),
neuper@37950
  3575
	       Calc ("Atools.is'_even",eval_is_even "#is_even_"),
neuper@38045
  3576
	       Calc ("Orderings.ord_class.less",eval_equ "#less_"),
neuper@37979
  3577
	       Thm ("not_false", num_str @{thm not_false}),
neuper@37979
  3578
	       Thm ("not_true", num_str @{thm not_true}),
neuper@38014
  3579
	       Calc ("Groups.plus_class.plus",eval_binop "#add_")
neuper@37950
  3580
	       ],
neuper@37979
  3581
      scr = EmptyScr
neuper@37950
  3582
      }:rls);
neuper@37950
  3583
(*.all powers over + distributed; atoms over * collected, other distributed
neuper@37950
  3584
   contains absolute minimum of thms for context in norm_Rational .*)
neuper@37950
  3585
val powers = prep_rls(
neuper@37950
  3586
  Rls {id = "powers", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3587
      erls = powers_erls, srls = Erls, calc = [], errpatts = [],
neuper@37969
  3588
      rules = [Thm ("realpow_multI", num_str @{thm realpow_multI}),
neuper@37950
  3589
	       (*"(r * s) ^^^ n = r ^^^ n * s ^^^ n"*)
neuper@37969
  3590
	       Thm ("realpow_pow",num_str @{thm realpow_pow}),
neuper@37950
  3591
	       (*"(a ^^^ b) ^^^ c = a ^^^ (b * c)"*)
neuper@37969
  3592
	       Thm ("realpow_oneI",num_str @{thm realpow_oneI}),
neuper@37950
  3593
	       (*"r ^^^ 1 = r"*)
neuper@37969
  3594
	       Thm ("realpow_minus_even",num_str @{thm realpow_minus_even}),
neuper@37950
  3595
	       (*"n is_even ==> (- r) ^^^ n = r ^^^ n" ?-->discard_minus?*)
neuper@37969
  3596
	       Thm ("realpow_minus_odd",num_str @{thm realpow_minus_odd}),
neuper@37950
  3597
	       (*"Not (n is_even) ==> (- r) ^^^ n = -1 * r ^^^ n"*)
neuper@37950
  3598
	       
neuper@37950
  3599
	       (*----- collect atoms over * -----*)
neuper@37969
  3600
	       Thm ("realpow_two_atom",num_str @{thm realpow_two_atom}),	
neuper@37950
  3601
	       (*"r is_atom ==> r * r = r ^^^ 2"*)
neuper@37969
  3602
	       Thm ("realpow_plus_1",num_str @{thm realpow_plus_1}),		
neuper@37950
  3603
	       (*"r is_atom ==> r * r ^^^ n = r ^^^ (n + 1)"*)
neuper@37969
  3604
	       Thm ("realpow_addI_atom",num_str @{thm realpow_addI_atom}),
neuper@37950
  3605
	       (*"r is_atom ==> r ^^^ n * r ^^^ m = r ^^^ (n + m)"*)
neuper@37950
  3606
neuper@37950
  3607
	       (*----- distribute none-atoms -----*)
neuper@37969
  3608
	       Thm ("realpow_def_atom",num_str @{thm realpow_def_atom}),
neuper@37950
  3609
	       (*"[| 1 < n; not(r is_atom) |]==>r ^^^ n = r * r ^^^ (n + -1)"*)
neuper@37969
  3610
	       Thm ("realpow_eq_oneI",num_str @{thm realpow_eq_oneI}),
neuper@37950
  3611
	       (*"1 ^^^ n = 1"*)
neuper@38014
  3612
	       Calc ("Groups.plus_class.plus",eval_binop "#add_")
neuper@37950
  3613
	       ],
neuper@37979
  3614
      scr = EmptyScr
neuper@37950
  3615
      }:rls);
neuper@37950
  3616
(*.contains absolute minimum of thms for context in norm_Rational.*)
neuper@37950
  3617
val rat_mult_divide = prep_rls(
neuper@37950
  3618
  Rls {id = "rat_mult_divide", preconds = [], 
neuper@37950
  3619
       rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3620
      erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@37969
  3621
      rules = [Thm ("rat_mult",num_str @{thm rat_mult}),
neuper@37950
  3622
	       (*(1)"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
neuper@37965
  3623
	       Thm ("times_divide_eq_right",num_str @{thm times_divide_eq_right}),
neuper@37950
  3624
	       (*(2)"?a * (?c / ?d) = ?a * ?c / ?d" must be [2],
neuper@37950
  3625
	       otherwise inv.to a / b / c = ...*)
neuper@37965
  3626
	       Thm ("times_divide_eq_left",num_str @{thm times_divide_eq_left}),
neuper@37950
  3627
	       (*"?a / ?b * ?c = ?a * ?c / ?b" order weights x^^^n too much
neuper@37950
  3628
		     and does not commute a / b * c ^^^ 2 !*)
neuper@37950
  3629
	       
neuper@37979
  3630
	       Thm ("divide_divide_eq_right", 
neuper@37979
  3631
                     num_str @{thm divide_divide_eq_right}),
neuper@37950
  3632
	       (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
neuper@37979
  3633
	       Thm ("divide_divide_eq_left",
neuper@37979
  3634
                     num_str @{thm divide_divide_eq_left}),
neuper@37950
  3635
	       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
neuper@48789
  3636
	       Calc ("Fields.inverse_class.divide"  ,eval_cancel "#divide_e")
neuper@37950
  3637
	       ],
neuper@37979
  3638
      scr = EmptyScr
neuper@37950
  3639
      }:rls);
neuper@37979
  3640
neuper@37950
  3641
(*.contains absolute minimum of thms for context in norm_Rational.*)
neuper@37950
  3642
val reduce_0_1_2 = prep_rls(
neuper@37950
  3643
  Rls{id = "reduce_0_1_2", preconds = [], rew_ord = ("dummy_ord", dummy_ord),
neuper@42451
  3644
      erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@37965
  3645
      rules = [(*Thm ("divide_1",num_str @{thm divide_1}),
neuper@37950
  3646
		 "?x / 1 = ?x" unnecess.for normalform*)
neuper@37965
  3647
	       Thm ("mult_1_left",num_str @{thm mult_1_left}),                 
neuper@37950
  3648
	       (*"1 * z = z"*)
neuper@37969
  3649
	       (*Thm ("real_mult_minus1",num_str @{thm real_mult_minus1}),
neuper@37950
  3650
	       "-1 * z = - z"*)
neuper@37969
  3651
	       (*Thm ("real_minus_mult_cancel",num_str @{thm real_minus_mult_cancel}),
neuper@37950
  3652
	       "- ?x * - ?y = ?x * ?y"*)
neuper@37950
  3653
neuper@37965
  3654
	       Thm ("mult_zero_left",num_str @{thm mult_zero_left}),        
neuper@37950
  3655
	       (*"0 * z = 0"*)
neuper@37965
  3656
	       Thm ("add_0_left",num_str @{thm add_0_left}),
neuper@37950
  3657
	       (*"0 + z = z"*)
neuper@37965
  3658
	       (*Thm ("right_minus",num_str @{thm right_minus}),
neuper@37950
  3659
	       "?z + - ?z = 0"*)
neuper@37950
  3660
neuper@37969
  3661
	       Thm ("sym_real_mult_2",
neuper@37969
  3662
                     num_str (@{thm real_mult_2} RS @{thm sym})),	
neuper@37950
  3663
	       (*"z1 + z1 = 2 * z1"*)
neuper@37969
  3664
	       Thm ("real_mult_2_assoc",num_str @{thm real_mult_2_assoc}),
neuper@37950
  3665
	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
neuper@37950
  3666
neuper@37965
  3667
	       Thm ("divide_zero_left",num_str @{thm divide_zero_left})
neuper@37950
  3668
	       (*"0 / ?x = 0"*)
neuper@37950
  3669
	       ], scr = EmptyScr}:rls);
neuper@37950
  3670
neuper@37950
  3671
(*erls for calculate_Rational; 
neuper@37950
  3672
  make local with FIXX@ME result:term *term list WN0609???SKMG*)
neuper@37950
  3673
val norm_rat_erls = prep_rls(
neuper@37950
  3674
  Rls {id = "norm_rat_erls", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3675
      erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@37950
  3676
      rules = [Calc ("Atools.is'_const",eval_const "#is_const_")
neuper@37979
  3677
	       ], scr = EmptyScr}:rls);
neuper@37979
  3678
neuper@37950
  3679
(*.consists of rls containing the absolute minimum of thms.*)
neuper@37950
  3680
(*040209: this version has been used by RL for his equations,
neuper@37950
  3681
which is now replaced by MGs version below
neuper@37950
  3682
vvv OLD VERSION !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!*)
neuper@37950
  3683
val norm_Rational = prep_rls(
neuper@37950
  3684
  Rls {id = "norm_Rational", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3685
      erls = norm_rat_erls, srls = Erls, calc = [], errpatts = [],
neuper@37950
  3686
      rules = [(*sequence given by operator precedence*)
neuper@37950
  3687
	       Rls_ discard_minus,
neuper@37950
  3688
	       Rls_ powers,
neuper@37950
  3689
	       Rls_ rat_mult_divide,
neuper@37950
  3690
	       Rls_ expand,
neuper@37950
  3691
	       Rls_ reduce_0_1_2,
neuper@37950
  3692
	       (*^^^^^^^^^ from RL -- not the latest one vvvvvvvvv*)
neuper@37950
  3693
	       Rls_ order_add_mult,
neuper@37950
  3694
	       Rls_ collect_numerals,
neuper@37950
  3695
	       Rls_ add_fractions_p,
neuper@37950
  3696
	       Rls_ cancel_p
neuper@37950
  3697
	       ],
neuper@37979
  3698
      scr = EmptyScr}:rls);
neuper@37979
  3699
neuper@37950
  3700
val norm_Rational_parenthesized = prep_rls(
neuper@37950
  3701
  Seq {id = "norm_Rational_parenthesized", preconds = []:term list, 
neuper@37950
  3702
       rew_ord = ("dummy_ord", dummy_ord),
neuper@37950
  3703
      erls = Atools_erls, srls = Erls,
neuper@42451
  3704
      calc = [], errpatts = [],
neuper@37950
  3705
      rules = [Rls_  norm_Rational, (*from RL -- not the latest one*)
neuper@37950
  3706
	       Rls_ discard_parentheses
neuper@37950
  3707
	       ],
neuper@37979
  3708
      scr = EmptyScr}:rls);      
neuper@37950
  3709
neuper@37950
  3710
(*WN030318???SK: simplifies all but cancel and common_nominator*)
neuper@37950
  3711
val simplify_rational = 
neuper@37950
  3712
    merge_rls "simplify_rational" expand_binoms
neuper@37950
  3713
    (append_rls "divide" calculate_Rational
neuper@37965
  3714
		[Thm ("divide_1",num_str @{thm divide_1}),
neuper@37950
  3715
		 (*"?x / 1 = ?x"*)
neuper@37978
  3716
		 Thm ("rat_mult",num_str @{thm rat_mult}),
neuper@37950
  3717
		 (*(1)"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
neuper@37965
  3718
		 Thm ("times_divide_eq_right",num_str @{thm times_divide_eq_right}),
neuper@37950
  3719
		 (*(2)"?a * (?c / ?d) = ?a * ?c / ?d" must be [2],
neuper@37950
  3720
		 otherwise inv.to a / b / c = ...*)
neuper@37965
  3721
		 Thm ("times_divide_eq_left",num_str @{thm times_divide_eq_left}),
neuper@37950
  3722
		 (*"?a / ?b * ?c = ?a * ?c / ?b"*)
neuper@37969
  3723
		 Thm ("add_minus",num_str @{thm add_minus}),
neuper@37950
  3724
		 (*"?a + ?b - ?b = ?a"*)
neuper@37969
  3725
		 Thm ("add_minus1",num_str @{thm add_minus1}),
neuper@37950
  3726
		 (*"?a - ?b + ?b = ?a"*)
neuper@37978
  3727
		 Thm ("divide_minus1",num_str @{thm divide_minus1})
neuper@37950
  3728
		 (*"?x / -1 = - ?x"*)
neuper@37950
  3729
(*
neuper@37950
  3730
,
neuper@37969
  3731
		 Thm ("",num_str @{thm })
neuper@37950
  3732
*)
neuper@37950
  3733
		 ]);
neuper@42451
  3734
*}
neuper@42451
  3735
ML {* 
neuper@37950
  3736
neuper@37950
  3737
(*---------vvv-------------MG ab 1.07.2003--------------vvv-----------*)
neuper@37950
  3738
neuper@37950
  3739
(* ------------------------------------------------------------------ *)
neuper@37950
  3740
(*                  Simplifier für beliebige Buchterme                *) 
neuper@37950
  3741
(* ------------------------------------------------------------------ *)
neuper@37950
  3742
(*----------------------- norm_Rational_mg ---------------------------*)
neuper@37950
  3743
(*. description of the simplifier see MG-DA.p.56ff .*)
neuper@37950
  3744
(* ------------------------------------------------------------------- *)
neuper@37950
  3745
val common_nominator_p_rls = prep_rls(
neuper@37950
  3746
  Rls {id = "common_nominator_p_rls", preconds = [],
neuper@37950
  3747
       rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3748
	 erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@37950
  3749
	 rules = 
neuper@37950
  3750
	 [Rls_ common_nominator_p
neuper@37950
  3751
	  (*FIXME.WN0401 ? redesign Rrls - use exhaustively on a term ?
neuper@37950
  3752
	    FIXME.WN0510 unnecessary nesting: introduce RRls_ : rls -> rule*)
neuper@37950
  3753
	  ], 
neuper@37950
  3754
	 scr = EmptyScr});
neuper@37950
  3755
(* ------------------------------------------------------------------- *)
neuper@37950
  3756
val cancel_p_rls = prep_rls(
neuper@37950
  3757
  Rls {id = "cancel_p_rls", preconds = [],
neuper@37950
  3758
       rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3759
	 erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@37950
  3760
	 rules = 
neuper@37950
  3761
	 [Rls_ cancel_p
neuper@37950
  3762
	  (*FIXME.WN.0401 ? redesign Rrls - use exhaustively on a term ?*)
neuper@37950
  3763
	  ], 
neuper@37950
  3764
	 scr = EmptyScr});
neuper@37950
  3765
(* -------------------------------------------------------------------- *)
neuper@37950
  3766
(*. makes 'normal' fractions; 'is_polyexp' inhibits double fractions;
neuper@37950
  3767
    used in initial part norm_Rational_mg, see example DA-M02-main.p.60.*)
neuper@37950
  3768
val rat_mult_poly = prep_rls(
neuper@37950
  3769
  Rls {id = "rat_mult_poly", preconds = [],
neuper@37950
  3770
       rew_ord = ("dummy_ord",dummy_ord), 
neuper@37950
  3771
	 erls =  append_rls "e_rls-is_polyexp" e_rls
neuper@37950
  3772
	         [Calc ("Poly.is'_polyexp", eval_is_polyexp "")], 
neuper@42451
  3773
	 srls = Erls, calc = [], errpatts = [],
neuper@37950
  3774
	 rules = 
neuper@37969
  3775
	 [Thm ("rat_mult_poly_l",num_str @{thm rat_mult_poly_l}),
neuper@37950
  3776
	  (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
neuper@37969
  3777
	  Thm ("rat_mult_poly_r",num_str @{thm rat_mult_poly_r})
neuper@37950
  3778
	  (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*)
neuper@37950
  3779
	  ], 
neuper@37950
  3780
	 scr = EmptyScr});
neuper@37979
  3781
neuper@37950
  3782
(* ------------------------------------------------------------------ *)
neuper@37950
  3783
(*. makes 'normal' fractions; 'is_polyexp' inhibits double fractions;
neuper@37950
  3784
    used in looping part norm_Rational_rls, see example DA-M02-main.p.60 
neuper@37950
  3785
    .. WHERE THE LATTER DOES ALWAYS WORK, BECAUSE erls = e_rls, 
neuper@37950
  3786
    I.E. THE RESPECTIVE ASSUMPTION IS STORED AND Thm APPLIED; WN051028 
neuper@37950
  3787
    ... WN0609???MG.*)
neuper@37950
  3788
val rat_mult_div_pow = prep_rls(
neuper@37950
  3789
  Rls {id = "rat_mult_div_pow", preconds = [], 
neuper@37950
  3790
       rew_ord = ("dummy_ord",dummy_ord), 
neuper@37950
  3791
       erls = e_rls,
neuper@37950
  3792
       (*FIXME.WN051028 append_rls "e_rls-is_polyexp" e_rls
neuper@37950
  3793
			[Calc ("Poly.is'_polyexp", eval_is_polyexp "")],
neuper@37950
  3794
         with this correction ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ we get 
neuper@37950
  3795
	 error "rational.sml.sml: diff.behav. in norm_Rational_mg 29" etc.
neuper@37950
  3796
         thus we decided to go on with this flaw*)
neuper@42451
  3797
		 srls = Erls, calc = [], errpatts = [],
neuper@37969
  3798
      rules = [Thm ("rat_mult",num_str @{thm rat_mult}),
neuper@37950
  3799
	       (*"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
neuper@37969
  3800
	       Thm ("rat_mult_poly_l",num_str @{thm rat_mult_poly_l}),
neuper@37950
  3801
	       (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
neuper@37969
  3802
	       Thm ("rat_mult_poly_r",num_str @{thm rat_mult_poly_r}),
neuper@37950
  3803
	       (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*)
neuper@37950
  3804
neuper@37979
  3805
	       Thm ("real_divide_divide1_mg",
neuper@37979
  3806
                     num_str @{thm real_divide_divide1_mg}),
neuper@37950
  3807
	       (*"y ~= 0 ==> (u / v) / (y / z) = (u * z) / (y * v)"*)
neuper@37979
  3808
	       Thm ("divide_divide_eq_right",
neuper@37979
  3809
                     num_str @{thm divide_divide_eq_right}),
neuper@37950
  3810
	       (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
neuper@37979
  3811
	       Thm ("divide_divide_eq_left",
neuper@37979
  3812
                     num_str @{thm divide_divide_eq_left}),
neuper@37950
  3813
	       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
neuper@48789
  3814
	       Calc ("Fields.inverse_class.divide"  ,eval_cancel "#divide_e"),
neuper@37950
  3815
	      
neuper@37969
  3816
	       Thm ("rat_power", num_str @{thm rat_power})
neuper@37950
  3817
		(*"(?a / ?b) ^^^ ?n = ?a ^^^ ?n / ?b ^^^ ?n"*)
neuper@37950
  3818
	       ],
neuper@37979
  3819
      scr = EmptyScr}:rls);
neuper@37950
  3820
(* ------------------------------------------------------------------ *)
neuper@37950
  3821
val rat_reduce_1 = prep_rls(
neuper@37950
  3822
  Rls {id = "rat_reduce_1", preconds = [], 
neuper@37950
  3823
       rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3824
       erls = e_rls, srls = Erls, calc = [], errpatts = [], 
neuper@37965
  3825
       rules = [Thm ("divide_1",num_str @{thm divide_1}),
neuper@37950
  3826
		(*"?x / 1 = ?x"*)
neuper@37965
  3827
		Thm ("mult_1_left",num_str @{thm mult_1_left})           
neuper@37950
  3828
		(*"1 * z = z"*)
neuper@37950
  3829
		],
neuper@37979
  3830
       scr = EmptyScr}:rls);
neuper@37950
  3831
(* ------------------------------------------------------------------ *)
neuper@37950
  3832
(*. looping part of norm_Rational(*_mg*) .*)
neuper@37950
  3833
val norm_Rational_rls = prep_rls(
neuper@37950
  3834
   Rls {id = "norm_Rational_rls", preconds = [], 
neuper@37950
  3835
       rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3836
       erls = norm_rat_erls, srls = Erls, calc = [], errpatts = [],
neuper@37950
  3837
       rules = [Rls_ common_nominator_p_rls,
neuper@37950
  3838
		Rls_ rat_mult_div_pow,
neuper@37950
  3839
		Rls_ make_rat_poly_with_parentheses,
neuper@37950
  3840
		Rls_ cancel_p_rls,(*FIXME:cancel_p does NOT order sometimes*)
neuper@37950
  3841
		Rls_ rat_reduce_1
neuper@37950
  3842
		],
neuper@37979
  3843
       scr = EmptyScr}:rls);
neuper@37950
  3844
(* ------------------------------------------------------------------ *)
neuper@37950
  3845
(*040109 'norm_Rational'(by RL) replaced by 'norm_Rational_mg'(MG)
neuper@37950
  3846
 just be renaming:*)
neuper@37950
  3847
val norm_Rational(*_mg*) = prep_rls(
neuper@37950
  3848
   Seq {id = "norm_Rational"(*_mg*), preconds = [], 
neuper@37950
  3849
       rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
  3850
       erls = norm_rat_erls, srls = Erls, calc = [], errpatts = [],
neuper@37980
  3851
       rules = [Rls_ discard_minus,
neuper@37950
  3852
		Rls_ rat_mult_poly,(* removes double fractions like a/b/c    *)
neuper@37950
  3853
		Rls_ make_rat_poly_with_parentheses, (*WN0510 also in(#)below*)
neuper@37950
  3854
		Rls_ cancel_p_rls, (*FIXME.MG:cancel_p does NOT order sometim*)
neuper@37950
  3855
		Rls_ norm_Rational_rls,   (* the main rls, looping (#)       *)
neuper@37979
  3856
		Rls_ discard_parentheses1 (* mult only                       *)
neuper@37950
  3857
		],
neuper@37979
  3858
       scr = EmptyScr}:rls);
neuper@37950
  3859
(* ------------------------------------------------------------------ *)
neuper@37950
  3860
neuper@42451
  3861
*}
neuper@42451
  3862
ML {* 
neuper@37967
  3863
ruleset' := overwritelthy @{theory} (!ruleset',
neuper@37950
  3864
  [("calculate_Rational", calculate_Rational),
neuper@37950
  3865
   ("calc_rat_erls",calc_rat_erls),
neuper@37950
  3866
   ("rational_erls", rational_erls),
neuper@37950
  3867
   ("cancel_p", cancel_p),
neuper@37950
  3868
   ("cancel", cancel),
neuper@37950
  3869
   ("common_nominator_p", common_nominator_p),
neuper@37950
  3870
   ("common_nominator_p_rls", common_nominator_p_rls),
neuper@37950
  3871
   ("common_nominator"  , common_nominator),
neuper@42407
  3872
   (*WN120410 ("discard_minus", discard_minus), is defined in Poly.thy*)
neuper@37950
  3873
   ("powers_erls", powers_erls),
neuper@37950
  3874
   ("powers", powers),
neuper@37950
  3875
   ("rat_mult_divide", rat_mult_divide),
neuper@37950
  3876
   ("reduce_0_1_2", reduce_0_1_2),
neuper@37950
  3877
   ("rat_reduce_1", rat_reduce_1),
neuper@37950
  3878
   ("norm_rat_erls", norm_rat_erls),
neuper@37950
  3879
   ("norm_Rational", norm_Rational),
neuper@37950
  3880
   ("norm_Rational_rls", norm_Rational_rls),
neuper@37950
  3881
   ("norm_Rational_parenthesized", norm_Rational_parenthesized),
neuper@37950
  3882
   ("rat_mult_poly", rat_mult_poly),
neuper@37950
  3883
   ("rat_mult_div_pow", rat_mult_div_pow),
neuper@37950
  3884
   ("cancel_p_rls", cancel_p_rls)
neuper@37950
  3885
   ]);
neuper@37950
  3886
neuper@37950
  3887
calclist':= overwritel (!calclist', 
neuper@37950
  3888
   [("is_expanded", ("Rational.is'_expanded", eval_is_expanded ""))
neuper@37950
  3889
    ]);
neuper@37950
  3890
neuper@37950
  3891
(** problems **)
neuper@37950
  3892
neuper@37950
  3893
store_pbt
neuper@37972
  3894
 (prep_pbt thy "pbl_simp_rat" [] e_pblID
neuper@37950
  3895
 (["rational","simplification"],
neuper@38083
  3896
  [("#Given" ,["Term t_t"]),
neuper@37979
  3897
   ("#Where" ,["t_t is_ratpolyexp"]),
neuper@37979
  3898
   ("#Find"  ,["normalform n_n"])
neuper@37950
  3899
  ],
neuper@37950
  3900
  append_rls "e_rls" e_rls [(*for preds in where_*)], 
neuper@38066
  3901
  SOME "Simplify t_t", 
neuper@37950
  3902
  [["simplification","of_rationals"]]));
neuper@37950
  3903
neuper@37950
  3904
(** methods **)
neuper@37950
  3905
neuper@37950
  3906
(*WN061025 this methods script is copied from (auto-generated) script
neuper@37950
  3907
  of norm_Rational in order to ease repair on inform*)
neuper@42439
  3908
store_met (prep_met thy "met_simp_rat" [] e_metID (["simplification","of_rationals"],
neuper@42439
  3909
	  [("#Given" ,["Term t_t"]),
neuper@42439
  3910
	   ("#Where" ,["t_t is_ratpolyexp"]),
neuper@42439
  3911
	   ("#Find"  ,["normalform n_n"])],
neuper@42439
  3912
	  {rew_ord'="tless_true", rls' = e_rls, calc = [], srls = e_rls, 
neuper@42439
  3913
	   prls = append_rls "simplification_of_rationals_prls" e_rls 
neuper@42439
  3914
				    [(*for preds in where_*) Calc ("Rational.is'_ratpolyexp", eval_is_ratpolyexp "")],
neuper@42516
  3915
				   crls = e_rls, errpats = [],
neuper@42439
  3916
   nrls = norm_Rational_rls},
neuper@42439
  3917
   "Script SimplifyScript (t_t::real) =                             " ^
neuper@42439
  3918
   "  ((Try (Rewrite_Set discard_minus False) @@                   " ^
neuper@42439
  3919
   "    Try (Rewrite_Set rat_mult_poly False) @@                    " ^
neuper@42439
  3920
   "    Try (Rewrite_Set make_rat_poly_with_parentheses False) @@   " ^
neuper@42439
  3921
   "    Try (Rewrite_Set cancel_p_rls False) @@                     " ^
neuper@42439
  3922
   "    (Repeat                                                     " ^
neuper@42439
  3923
   "     ((Try (Rewrite_Set common_nominator_p_rls False) @@        " ^
neuper@42439
  3924
   "       Try (Rewrite_Set rat_mult_div_pow False) @@              " ^
neuper@42439
  3925
   "       Try (Rewrite_Set make_rat_poly_with_parentheses False) @@" ^
neuper@42439
  3926
   "       Try (Rewrite_Set cancel_p_rls False) @@                  " ^
neuper@42439
  3927
   "       Try (Rewrite_Set rat_reduce_1 False)))) @@               " ^
neuper@42439
  3928
   "    Try (Rewrite_Set discard_parentheses1 False))               " ^
neuper@42439
  3929
   "   t_t)"));
neuper@37979
  3930
neuper@37950
  3931
*}
neuper@48880
  3932
ML {*"test"*}
neuper@42451
  3933
end (* theory*)