src/Tools/isac/Knowledge/Rational.thy
author wneuper <walther.neuper@jku.at>
Mon, 19 Jul 2021 15:34:54 +0200
changeset 60335 7701598a2182
parent 60331 40eb8aa2b0d6
child 60337 cbad4e18e91b
permissions -rw-r--r--
ALL const_name replaces (others cannot be replaced)
     1 (* rationals, fractions of multivariate polynomials over the real field
     2    author: isac team
     3    Copyright (c) isac team 2002, 2013
     4    Use is subject to license terms.
     5 
     6    depends on Poly (and not on Atools), because 
     7    fractions with _normalised_ polynomials are canceled, added, etc.
     8 *)
     9 
    10 theory Rational 
    11 imports Poly GCD_Poly_ML
    12 begin
    13 
    14 section \<open>Constants for evaluation by "Rule.Eval"\<close>
    15 consts
    16 
    17   is_expanded    :: "real => bool" ("_ is'_expanded")     (*RL->Poly.thy*)
    18   is_ratpolyexp  :: "real => bool" ("_ is'_ratpolyexp") 
    19   get_denominator :: "real => real"
    20   get_numerator   :: "real => real"           
    21 
    22 ML \<open>
    23 (*.the expression contains + - * ^ / only ?.*)
    24 fun is_ratpolyexp (Free _) = true
    25   | is_ratpolyexp (Const (\<^const_name>\<open>plus\<close>,_) $ Free _ $ Free _) = true
    26   | is_ratpolyexp (Const (\<^const_name>\<open>minus\<close>,_) $ Free _ $ Free _) = true
    27   | is_ratpolyexp (Const (\<^const_name>\<open>times\<close>,_) $ Free _ $ Free _) = true
    28   | is_ratpolyexp (Const (\<^const_name>\<open>powr\<close>,_) $ Free _ $ Free _) = true
    29   | is_ratpolyexp (Const (\<^const_name>\<open>divide\<close>,_) $ Free _ $ Free _) = true
    30   | is_ratpolyexp (Const (\<^const_name>\<open>plus\<close>,_) $ t1 $ t2) = 
    31                ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
    32   | is_ratpolyexp (Const (\<^const_name>\<open>minus\<close>,_) $ t1 $ t2) = 
    33                ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
    34   | is_ratpolyexp (Const (\<^const_name>\<open>times\<close>,_) $ t1 $ t2) = 
    35                ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
    36   | is_ratpolyexp (Const (\<^const_name>\<open>powr\<close>,_) $ t1 $ t2) = 
    37                ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
    38   | is_ratpolyexp (Const (\<^const_name>\<open>divide\<close>,_) $ t1 $ t2) = 
    39                ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
    40   | is_ratpolyexp _ = false;
    41 
    42 (*("is_ratpolyexp", ("Rational.is_ratpolyexp", eval_is_ratpolyexp ""))*)
    43 fun eval_is_ratpolyexp (thmid:string) _ 
    44 		       (t as (Const (\<^const_name>\<open>Rational.is_ratpolyexp\<close>, _) $ arg)) thy =
    45     if is_ratpolyexp arg
    46     then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
    47 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
    48     else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
    49 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
    50   | eval_is_ratpolyexp _ _ _ _ = NONE; 
    51 
    52 (*("get_denominator", ("Rational.get_denominator", eval_get_denominator ""))*)
    53 fun eval_get_denominator (thmid:string) _ 
    54 		      (t as Const (\<^const_name>\<open>Rational.get_denominator\<close>, _) $
    55               (Const (\<^const_name>\<open>divide\<close>, _) $ _(*num*) $
    56                 denom)) thy = 
    57       SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy denom) "", 
    58 	            HOLogic.Trueprop $ (TermC.mk_equality (t, denom)))
    59   | eval_get_denominator _ _ _ _ = NONE; 
    60 
    61 (*("get_numerator", ("Rational.get_numerator", eval_get_numerator ""))*)
    62 fun eval_get_numerator (thmid:string) _ 
    63       (t as Const (\<^const_name>\<open>Rational.get_numerator\<close>, _) $
    64           (Const (\<^const_name>\<open>divide\<close>, _) $num
    65             $denom )) thy = 
    66     SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy num) "", 
    67 	    HOLogic.Trueprop $ (TermC.mk_equality (t, num)))
    68   | eval_get_numerator _ _ _ _ = NONE; 
    69 \<close>
    70 
    71 section \<open>Theorems for rewriting\<close>
    72 
    73 axiomatization (* naming due to Isabelle2002, but not contained in Isabelle2002; 
    74                   many thms are due to RL and can be removed with updating the equation solver;
    75                   TODO: replace by equivalent thms in recent Isabelle201x *) 
    76 where
    77   mult_cross:      "[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)" and
    78   mult_cross1:     "   b ~= 0            ==> (a / b = c    ) = (a     = b * c)" and
    79   mult_cross2:     "           d ~= 0    ==> (a     = c / d) = (a * d =     c)" and
    80                   
    81   add_minus:       "a + b - b = a"(*RL->Poly.thy*) and
    82   add_minus1:      "a - b + b = a"(*RL->Poly.thy*) and
    83                   
    84   rat_mult:        "a / b * (c / d) = a * c / (b * d)"(*?Isa02*)  and
    85   rat_mult2:       "a / b *  c      = a * c /  b     "(*?Isa02*) and
    86 
    87   rat_mult_poly_l: "c is_polyexp ==> c * (a / b) = c * a /  b" and
    88   rat_mult_poly_r: "c is_polyexp ==> (a / b) * c = a * c /  b" and
    89 
    90 (*real_times_divide1_eq .. Isa02*) 
    91   real_times_divide_1_eq:  "-1 * (c / d) = -1 * c / d " and
    92   real_times_divide_num:   "a is_const ==> a * (c / d) = a * c / d " and
    93 
    94   real_mult_div_cancel2:   "k ~= 0 ==> m * k / (n * k) = m / n" and
    95 (*real_mult_div_cancel1:   "k ~= 0 ==> k * m / (k * n) = m / n"..Isa02*)
    96 			  
    97   real_divide_divide1:     "y ~= 0 ==> (u / v) / (y / z) = (u / v) * (z / y)" and
    98   real_divide_divide1_mg:  "y ~= 0 ==> (u / v) / (y / z) = (u * z) / (y * v)" and
    99 (*real_divide_divide2_eq:  "x / y / z = x / (y * z)"..Isa02*)
   100 			  
   101   rat_power:               "(a / b) \<up> n = (a \<up> n) / (b \<up> n)" and
   102 
   103   rat_add:         "[| a is_const; b is_const; c is_const; d is_const |] ==> 
   104 	           a / c + b / d = (a * d + b * c) / (c * d)" and
   105   rat_add_assoc:   "[| a is_const; b is_const; c is_const; d is_const |] ==> 
   106 	           a / c +(b / d + e) = (a * d + b * c)/(d * c) + e" and
   107   rat_add1:        "[| a is_const; b is_const; c is_const |] ==> 
   108 	           a / c + b / c = (a + b) / c" and
   109   rat_add1_assoc:   "[| a is_const; b is_const; c is_const |] ==> 
   110 	           a / c + (b / c + e) = (a + b) / c + e" and
   111   rat_add2:        "[| a is_const; b is_const; c is_const |] ==> 
   112 	           a / c + b = (a + b * c) / c" and
   113   rat_add2_assoc:  "[| a is_const; b is_const; c is_const |] ==> 
   114 	           a / c + (b + e) = (a + b * c) / c + e" and
   115   rat_add3:        "[| a is_const; b is_const; c is_const |] ==> 
   116 	           a + b / c = (a * c + b) / c" and
   117   rat_add3_assoc:   "[| a is_const; b is_const; c is_const |] ==> 
   118 	           a + (b / c + e) = (a * c + b) / c + e"
   119 
   120 section \<open>Cancellation and addition of fractions\<close>
   121 subsection \<open>Conversion term <--> poly\<close>
   122 subsubsection \<open>Convert a term to the internal representation of a multivariate polynomial\<close>
   123 ML \<open>
   124 fun monom_of_term vs (c, es) (t as Const _) =
   125     (c, list_update es (find_index (curry op = t) vs) 1)
   126   | monom_of_term _ (c, es) (t as (Const (\<^const_name>\<open>numeral\<close>, _) $ _)) =
   127     (t |> HOLogic.dest_number |> snd |> curry op * c, es) (*several numerals in one monom*)
   128   | monom_of_term _ (c, es) (t as (Const (\<^const_name>\<open>uminus\<close>, _) $ _)) =
   129     (t |> HOLogic.dest_number |> snd |> curry op * c, es) (*several numerals in one monom*)
   130   | monom_of_term  vs (c, es) (t as Free _) =
   131     (c, list_update es (find_index (curry op = t) vs) 1)
   132   | monom_of_term  vs (c, es) (Const (\<^const_name>\<open>Transcendental.powr\<close>, _) $ (b as Free _) $
   133       (e as Const (\<^const_name>\<open>numeral\<close>, _) $ _)) =
   134     (c, list_update es (find_index (curry op = b) vs) (e |> HOLogic.dest_number |> snd))
   135   | monom_of_term  vs (c, es) (Const (\<^const_name>\<open>Transcendental.powr\<close>, _) $ (b as Free _) $
   136       (e as Const (\<^const_name>\<open>uminus\<close>, _) $ _)) =
   137     (c, list_update es (find_index (curry op = b) vs) (e |> HOLogic.dest_number |> snd))
   138 
   139   | monom_of_term vs (c, es) (Const (\<^const_name>\<open>times\<close>, _) $ m1 $ m2) =
   140     let val (c', es') = monom_of_term vs (c, es) m1
   141     in monom_of_term vs (c', es') m2 end
   142   | monom_of_term _ _ t = raise ERROR ("poly malformed 1 with " ^ UnparseC.term t)
   143 
   144 (*-------v------*)
   145 fun monoms_of_term vs (t as Const _) =
   146     [monom_of_term  vs (1, replicate (length vs) 0) t]
   147   | monoms_of_term vs (t as Const (\<^const_name>\<open>numeral\<close>, _) $ _) =
   148     [monom_of_term  vs (1, replicate (length vs) 0) t]
   149   | monoms_of_term vs (t as Const (\<^const_name>\<open>uminus\<close>, _) $ _) =
   150     [monom_of_term  vs (1, replicate (length vs) 0) t]
   151   | monoms_of_term vs (t as Free _) =
   152     [monom_of_term  vs (1, replicate (length vs) 0) t]
   153   | monoms_of_term vs (t as Const (\<^const_name>\<open>powr\<close>, _) $ _ $  _) =
   154     [monom_of_term  vs (1, replicate (length vs) 0) t]
   155   | monoms_of_term vs (t as Const (\<^const_name>\<open>times\<close>, _) $ _ $  _) =
   156     [monom_of_term  vs (1, replicate (length vs) 0) t]
   157   | monoms_of_term vs (Const (\<^const_name>\<open>plus\<close>, _) $ ms1 $ ms2) =
   158     (monoms_of_term vs ms1) @ (monoms_of_term vs ms2)
   159   | monoms_of_term _ t = raise ERROR ("poly malformed 2 with " ^ UnparseC.term t)
   160 
   161 (* convert a term to the internal representation of a multivariate polynomial;
   162   the conversion is quite liberal, see test --- fun poly_of_term ---:
   163 * the order of variables and the parentheses within a monomial are arbitrary
   164 * the coefficient may be somewhere
   165 * he order and the parentheses within monomials are arbitrary
   166 But the term must be completely expand + over * (laws of distributivity are not applicable).
   167 
   168 The function requires the free variables as strings already given, 
   169 because the gcd involves 2 polynomials (with the same length for their list of exponents).
   170 *)
   171 fun poly_of_term vs (t as Const (\<^const_name>\<open>plus\<close>, _) $ _ $ _) =
   172     (SOME (t |> monoms_of_term vs |> order)
   173       handle ERROR _ => NONE)
   174   | poly_of_term vs t =
   175     (SOME [monom_of_term vs (1, replicate (length vs) 0) t]
   176       handle ERROR _ => NONE)
   177 
   178 fun is_poly t =
   179   let
   180     val vs = TermC.vars_of t
   181   in 
   182     case poly_of_term vs t of SOME _ => true | NONE => false
   183   end
   184 val is_expanded = is_poly   (* TODO: check names *)
   185 val is_polynomial = is_poly (* TODO: check names *)
   186 \<close>
   187 
   188 subsubsection \<open>Convert internal representation of a multivariate polynomial to a term\<close>
   189 ML \<open>
   190 fun term_of_es _ _ _ [] = [] (*assumes same length for vs and es*)
   191   | term_of_es baseT expT (_ :: vs) (0 :: es) = [] @ term_of_es baseT expT vs es
   192   | term_of_es baseT expT (v :: vs) (1 :: es) = v :: term_of_es baseT expT vs es
   193   | term_of_es baseT expT (v :: vs) (e :: es) =
   194     Const (\<^const_name>\<open>Transcendental.powr\<close>, [baseT, expT] ---> baseT) $ v $ (HOLogic.mk_number expT e)
   195     :: term_of_es baseT expT vs es
   196   | term_of_es _ _ _ _ = raise ERROR "term_of_es: length vs <> length es"
   197 
   198 fun term_of_monom baseT expT vs ((c, es): monom) =
   199     let val es' = term_of_es baseT expT vs es
   200     in 
   201       if c = 1 
   202       then 
   203         if es' = [] (*if es = [0,0,0,...]*)
   204         then HOLogic.mk_number baseT c
   205         else foldl (HOLogic.mk_binop "Groups.times_class.times") (hd es', tl es')
   206       else foldl (HOLogic.mk_binop "Groups.times_class.times")
   207         (HOLogic.mk_number baseT c, es') 
   208     end
   209 
   210 fun term_of_poly baseT expT vs p =
   211   let val monos = map (term_of_monom baseT expT vs) p
   212   in foldl (HOLogic.mk_binop \<^const_name>\<open>plus\<close>) (hd monos, tl monos) end
   213 \<close>
   214 
   215 subsection \<open>Apply gcd_poly for cancelling and adding fractions as terms\<close>
   216 ML \<open>
   217 fun mk_noteq_0 baseT t = 
   218   Const (\<^const_name>\<open>Not\<close>, HOLogic.boolT --> HOLogic.boolT) $ 
   219     (Const (\<^const_name>\<open>HOL.eq\<close>, [baseT, baseT] ---> HOLogic.boolT) $ t $ HOLogic.mk_number HOLogic.realT 0)
   220 
   221 fun mk_asms baseT ts =
   222   let val as' = filter_out TermC.is_num ts (* asm like "2 ~= 0" is needless *)
   223   in map (mk_noteq_0 baseT) as' end
   224 \<close>
   225 
   226 subsubsection \<open>Factor out gcd for cancellation\<close>
   227 ML \<open>
   228 fun check_fraction t =
   229   case t of
   230     Const (\<^const_name>\<open>divide\<close>, _) $ numerator $ denominator
   231       => SOME (numerator, denominator)
   232   | _ => NONE
   233 
   234 (* prepare a term for cancellation by factoring out the gcd
   235   assumes: is a fraction with outmost "/"*)
   236 fun factout_p_ (thy: theory) t =
   237   let val opt = check_fraction t
   238   in
   239     case opt of 
   240       NONE => NONE
   241     | SOME (numerator, denominator) =>
   242       let
   243         val vs = TermC.vars_of t
   244         val baseT = type_of numerator
   245         val expT = HOLogic.realT
   246       in
   247         case (poly_of_term vs numerator, poly_of_term vs denominator) of
   248           (SOME a, SOME b) =>
   249             let
   250               val ((a', b'), c) = gcd_poly a b
   251               val es = replicate (length vs) 0 
   252             in
   253               if c = [(1, es)] orelse c = [(~1, es)]
   254               then NONE
   255               else 
   256                 let
   257                   val b't = term_of_poly baseT expT vs b'
   258                   val ct = term_of_poly baseT expT vs c
   259                   val t' = 
   260                     HOLogic.mk_binop \<^const_name>\<open>divide\<close> 
   261                       (HOLogic.mk_binop \<^const_name>\<open>times\<close>
   262                         (term_of_poly baseT expT vs a', ct),
   263                        HOLogic.mk_binop \<^const_name>\<open>times\<close> (b't, ct))
   264                 in SOME (t', mk_asms baseT [b't, ct]) end
   265             end
   266         | _ => NONE : (term * term list) option
   267       end
   268   end
   269 \<close>
   270 
   271 subsubsection \<open>Cancel a fraction\<close>
   272 ML \<open>
   273 (* cancel a term by the gcd ("" denote terms with internal algebraic structure)
   274   cancel_p_ :: theory \<Rightarrow> term  \<Rightarrow> (term \<times> term list) option
   275   cancel_p_ thy "a / b" = SOME ("a' / b'", ["b' \<noteq> 0"])
   276   assumes: a is_polynomial  \<and>  b is_polynomial  \<and>  b \<noteq> 0
   277   yields
   278     SOME ("a' / b'", ["b' \<noteq> 0"]). gcd_poly a b \<noteq> 1  \<and>  gcd_poly a b \<noteq> -1  \<and>  
   279       a' * gcd_poly a b = a  \<and>  b' * gcd_poly a b = b
   280     \<or> NONE *)
   281 fun cancel_p_ (_: theory) t =
   282   let val opt = check_fraction t
   283   in
   284     case opt of 
   285       NONE => NONE
   286     | SOME (numerator, denominator) =>
   287       let
   288         val vs = TermC.vars_of t
   289         val baseT = type_of numerator
   290         val expT = HOLogic.realT
   291       in
   292         case (poly_of_term vs numerator, poly_of_term vs denominator) of
   293           (SOME a, SOME b) =>
   294             let
   295               val ((a', b'), c) = gcd_poly a b
   296               val es = replicate (length vs) 0 
   297             in
   298               if c = [(1, es)] orelse c = [(~1, es)]
   299               then NONE
   300               else 
   301                 let
   302                   val bt' = term_of_poly baseT expT vs b'
   303                   val ct = term_of_poly baseT expT vs c
   304                   val t' = 
   305                     HOLogic.mk_binop \<^const_name>\<open>divide\<close> 
   306                       (term_of_poly baseT expT vs a', bt')
   307                   val asm = mk_asms baseT [bt']
   308                 in SOME (t', asm) end
   309             end
   310         | _ => NONE : (term * term list) option
   311       end
   312   end
   313 \<close>
   314 
   315 subsubsection \<open>Factor out to a common denominator for addition\<close>
   316 ML \<open>
   317 (* addition of fractions allows (at most) one non-fraction (a monomial) *)
   318 fun check_frac_sum 
   319     (Const (\<^const_name>\<open>plus\<close>, _) $ 
   320       (Const (\<^const_name>\<open>divide\<close>, _) $ n1 $ d1) $
   321       (Const (\<^const_name>\<open>divide\<close>, _) $ n2 $ d2))
   322     = SOME ((n1, d1), (n2, d2))
   323   | check_frac_sum 
   324     (Const (\<^const_name>\<open>plus\<close>, _) $ 
   325       nofrac $ 
   326       (Const (\<^const_name>\<open>divide\<close>, _) $ n2 $ d2))
   327     = SOME ((nofrac, Free ("1", HOLogic.realT)), (n2, d2))
   328   | check_frac_sum 
   329     (Const (\<^const_name>\<open>plus\<close>, _) $ 
   330       (Const (\<^const_name>\<open>divide\<close>, _) $ n1 $ d1) $ 
   331       nofrac)
   332     = SOME ((n1, d1), (nofrac, HOLogic.mk_number HOLogic.realT 1))
   333   | check_frac_sum _ = NONE  
   334 
   335 (* prepare a term for addition by providing the least common denominator as a product
   336   assumes: is a term with outmost "+" and at least one outmost "/" in respective summands*)
   337 fun common_nominator_p_ (_: theory) t =
   338   let val opt = check_frac_sum t
   339   in
   340     case opt of 
   341       NONE => NONE
   342     | SOME ((n1, d1), (n2, d2)) =>
   343       let
   344         val vs = TermC.vars_of t
   345       in
   346         case (poly_of_term vs d1, poly_of_term vs d2) of
   347           (SOME a, SOME b) =>
   348             let
   349               val ((a', b'), c) = gcd_poly a b
   350               val (baseT, expT) = (type_of n1, HOLogic.realT)
   351               val [d1', d2', c'] = map (term_of_poly baseT expT vs) [a', b', c]
   352               (*----- minimum of parentheses & nice result, but breaks tests: -------------
   353               val denom = HOLogic.mk_binop \<^const_name>\<open>times\<close> 
   354                 (HOLogic.mk_binop \<^const_name>\<open>times\<close> (d1', d2'), c') -------------*)
   355               val denom =
   356                 if c = [(1, replicate (length vs) 0)]
   357                 then HOLogic.mk_binop \<^const_name>\<open>times\<close> (d1', d2')
   358                 else
   359                   HOLogic.mk_binop \<^const_name>\<open>times\<close> (c',
   360                   HOLogic.mk_binop \<^const_name>\<open>times\<close> (d1', d2')) (*--------------*)
   361               val t' =
   362                 HOLogic.mk_binop \<^const_name>\<open>plus\<close>
   363                   (HOLogic.mk_binop \<^const_name>\<open>divide\<close>
   364                     (HOLogic.mk_binop \<^const_name>\<open>times\<close> (n1, d2'), denom),
   365                   HOLogic.mk_binop \<^const_name>\<open>divide\<close> 
   366                     (HOLogic.mk_binop \<^const_name>\<open>times\<close> (n2, d1'), denom))
   367               val asm = mk_asms baseT [d1', d2', c']
   368             in SOME (t', asm) end
   369         | _ => NONE : (term * term list) option
   370       end
   371   end
   372 \<close>
   373 
   374 subsubsection \<open>Addition of at least one fraction within a sum\<close>
   375 ML \<open>
   376 (* add fractions
   377   assumes: is a term with outmost "+" and at least one outmost "/" in respective summands
   378   NOTE: the case "(_ + _) + _" need not be considered due to iterated addition.*)
   379 fun add_fraction_p_ (_: theory) t =
   380   case check_frac_sum t of 
   381     NONE => NONE
   382   | SOME ((n1, d1), (n2, d2)) =>
   383     let
   384       val vs = TermC.vars_of t
   385     in
   386       case (poly_of_term vs n1, poly_of_term vs d1, poly_of_term vs n2, poly_of_term vs d2) of
   387         (SOME _, SOME a, SOME _, SOME b) =>
   388           let
   389             val ((a', b'), c) = gcd_poly a b
   390             val (baseT, expT) = (type_of n1, HOLogic.realT)
   391             val nomin = term_of_poly baseT expT vs 
   392               (((the (poly_of_term vs n1)) %%*%% b') %%+%% ((the (poly_of_term vs n2)) %%*%% a')) 
   393             val denom = term_of_poly baseT expT vs ((c %%*%% a') %%*%% b')
   394             val t' = HOLogic.mk_binop \<^const_name>\<open>divide\<close> (nomin, denom)
   395           in SOME (t', mk_asms baseT [denom]) end
   396       | _ => NONE : (term * term list) option
   397     end
   398 \<close>
   399 
   400 section \<open>Embed cancellation and addition into rewriting\<close>
   401 
   402 subsection \<open>Rulesets and predicate for embedding\<close>
   403 ML \<open>
   404 (* evaluates conditions in calculate_Rational *)
   405 val calc_rat_erls =
   406   prep_rls'
   407     (Rule_Def.Repeat {id = "calc_rat_erls", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
   408       erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   409       rules = 
   410        [Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches "#matches_"),
   411         Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
   412         Rule.Eval ("Prog_Expr.is_const", Prog_Expr.eval_const "#is_const_"),
   413         Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
   414         Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false})], 
   415       scr = Rule.Empty_Prog});
   416 
   417 (* simplifies expressions with numerals;
   418    does NOT rearrange the term by AC-rewriting; thus terms with variables 
   419    need to have constants to be commuted together respectively           *)
   420 val calculate_Rational =
   421   prep_rls' (Rule_Set.merge "calculate_Rational"
   422     (Rule_Def.Repeat {id = "divide", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
   423       erls = calc_rat_erls, srls = Rule_Set.Empty,
   424       calc = [], errpatts = [],
   425       rules = 
   426         [\<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
   427 
   428         Rule.Thm ("minus_divide_left", ThmC.numerals_to_Free (@{thm minus_divide_left} RS @{thm sym})),
   429           (*SYM - ?x / ?y = - (?x / ?y)  may come from subst*)
   430         \<^rule_thm>\<open>rat_add\<close>,
   431           (*"[| a is_const; b is_const; c is_const; d is_const |] ==> \
   432           \a / c + b / d = (a * d) / (c * d) + (b * c ) / (d * c)"*)
   433         \<^rule_thm>\<open>rat_add1\<close>,
   434           (*"[| a is_const; b is_const; c is_const |] ==> a / c + b / c = (a + b) / c"*)
   435         \<^rule_thm>\<open>rat_add2\<close>,
   436           (*"[| ?a is_const; ?b is_const; ?c is_const |] ==> ?a / ?c + ?b = (?a + ?b * ?c) / ?c"*)
   437         \<^rule_thm>\<open>rat_add3\<close>,
   438           (*"[| a is_const; b is_const; c is_const |] ==> a + b / c = (a * c) / c + b / c"\
   439           .... is_const to be omitted here FIXME*)
   440         
   441         \<^rule_thm>\<open>rat_mult\<close>, 
   442           (*a / b * (c / d) = a * c / (b * d)*)
   443         \<^rule_thm>\<open>times_divide_eq_right\<close>,
   444           (*?x * (?y / ?z) = ?x * ?y / ?z*)
   445         \<^rule_thm>\<open>times_divide_eq_left\<close>,
   446           (*?y / ?z * ?x = ?y * ?x / ?z*)
   447         
   448         \<^rule_thm>\<open>real_divide_divide1\<close>,
   449           (*"?y ~= 0 ==> ?u / ?v / (?y / ?z) = ?u / ?v * (?z / ?y)"*)
   450         \<^rule_thm>\<open>divide_divide_eq_left\<close>,
   451           (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
   452         
   453         \<^rule_thm>\<open>rat_power\<close>,
   454           (*"(?a / ?b)  \<up>  ?n = ?a  \<up>  ?n / ?b  \<up>  ?n"*)
   455         
   456         \<^rule_thm>\<open>mult_cross\<close>,
   457           (*"[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)*)
   458         \<^rule_thm>\<open>mult_cross1\<close>,
   459           (*"   b ~= 0            ==> (a / b = c    ) = (a     = b * c)*)
   460         \<^rule_thm>\<open>mult_cross2\<close>
   461           (*"           d ~= 0    ==> (a     = c / d) = (a * d =     c)*)], 
   462       scr = Rule.Empty_Prog})
   463     calculate_Poly);
   464 
   465 (*("is_expanded", ("Rational.is_expanded", eval_is_expanded ""))*)
   466 fun eval_is_expanded (thmid:string) _ 
   467 		       (t as (Const (\<^const_name>\<open>Rational.is_expanded\<close>, _) $ arg)) thy = 
   468     if is_expanded arg
   469     then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
   470 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
   471     else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
   472 	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
   473   | eval_is_expanded _ _ _ _ = NONE;
   474 \<close>
   475 calculation is_expanded = \<open>eval_is_expanded ""\<close>
   476 ML \<open>
   477 val rational_erls = 
   478   Rule_Set.merge "rational_erls" calculate_Rational 
   479     (Rule_Set.append_rules "is_expanded" Atools_erls 
   480       [\<^rule_eval>\<open>is_expanded\<close> (eval_is_expanded "")]);
   481 \<close>
   482 
   483 subsection \<open>Embed cancellation into rewriting\<close>
   484 ML \<open>
   485 (**)local (* cancel_p *)
   486 
   487 val {rules = rules, rew_ord = (_, ro), ...} = Rule_Set.rep (assoc_rls' @{theory} "rev_rew_p");
   488 
   489 fun init_state thy eval_rls ro t =
   490   let
   491     val SOME (t', _) = factout_p_ thy t;
   492     val SOME (t'', asm) = cancel_p_ thy t;
   493     val der = Derive.steps_reverse thy eval_rls rules ro NONE t';
   494     val der = der @ 
   495       [(\<^rule_thm>\<open>real_mult_div_cancel2\<close>, (t'', asm))]
   496     val rs = (Rule.distinct' o (map #1)) der
   497   	val rs = filter_out (ThmC.member'
   498   	  ["sym_real_add_zero_left", "sym_real_mult_0", "sym_real_mult_1"]) rs
   499   in (t, t'', [rs(*one in order to ease locate_rule*)], der) end;
   500 
   501 fun locate_rule thy eval_rls ro [rs] t r =
   502     if member op = ((map (Rule.thm_id)) rs) (Rule.thm_id r)
   503     then 
   504       let val ropt = Rewrite.rewrite_ thy ro eval_rls true (Rule.thm r) t;
   505       in
   506         case ropt of SOME ta => [(r, ta)]
   507 	      | NONE => ((*tracing 
   508 	          ("### locate_rule:  rewrite " ^ Rule.thm_id r ^ " " ^ UnparseC.term t ^ " = NONE");*) []) 
   509 			end
   510     else ((*tracing ("### locate_rule:  " ^ Rule.thm_id r ^ " not mem rrls");*) [])
   511   | locate_rule _ _ _ _ _ _ = raise ERROR "locate_rule: doesnt match rev-sets in istate";
   512 
   513 fun next_rule thy eval_rls ro [rs] t =
   514     let
   515       val der = Derive.do_one thy eval_rls rs ro NONE t;
   516     in case der of (_, r, _) :: _ => SOME r | _ => NONE end
   517   | next_rule _ _ _ _ _ = raise ERROR ("next_rule: doesnt match rev-sets in istate");
   518 
   519 fun attach_form (_: Rule.rule list list) (_: term) (_: term) = 
   520   [(*TODO*)]: ( Rule.rule * (term * term list)) list;
   521 
   522 (**)in(**)
   523 
   524 val cancel_p = 
   525   Rule_Set.Rrls {id = "cancel_p", prepat = [],
   526 	rew_ord=("ord_make_polynomial", ord_make_polynomial false \<^theory>),
   527 	erls = rational_erls, 
   528 	calc = 
   529 	  [("PLUS", (\<^const_name>\<open>plus\<close>, (**)eval_binop "#add_")),
   530 	  ("TIMES" , (\<^const_name>\<open>times\<close>, (**)eval_binop "#mult_")),
   531 	  ("DIVIDE", (\<^const_name>\<open>divide\<close>, Prog_Expr.eval_cancel "#divide_e")),
   532 	  ("POWER", (\<^const_name>\<open>powr\<close>, (**)eval_binop "#power_"))],
   533     errpatts = [],
   534 	scr =
   535 	  Rule.Rfuns {init_state  = init_state \<^theory> Atools_erls ro,
   536 		normal_form = cancel_p_ \<^theory>, 
   537 		locate_rule = locate_rule \<^theory> Atools_erls ro,
   538 		next_rule   = next_rule \<^theory> Atools_erls ro,
   539 		attach_form = attach_form}}
   540 (**)end(* local cancel_p *)
   541 \<close>
   542 
   543 subsection \<open>Embed addition into rewriting\<close>
   544 ML \<open>
   545 (**)local (* add_fractions_p *)
   546 
   547 (*val {rules = rules, rew_ord = (_, ro), ...} = Rule_Set.rep (assoc_rls "make_polynomial");*)
   548 val {rules, rew_ord=(_,ro),...} = Rule_Set.rep (assoc_rls' @{theory} "rev_rew_p");
   549 
   550 fun init_state thy eval_rls ro t =
   551   let 
   552     val SOME (t',_) = common_nominator_p_ thy t;
   553     val SOME (t'', asm) = add_fraction_p_ thy t;
   554     val der = Derive.steps_reverse thy eval_rls rules ro NONE t';
   555     val der = der @ 
   556       [(\<^rule_thm>\<open>real_mult_div_cancel2\<close>, (t'',asm))]
   557     val rs = (Rule.distinct' o (map #1)) der;
   558     val rs = filter_out (ThmC.member'
   559       ["sym_real_add_zero_left", "sym_real_mult_0", "sym_real_mult_1"]) rs;
   560   in (t, t'', [rs(*here only _ONE_*)], der) end;
   561 
   562 fun locate_rule thy eval_rls ro [rs] t r =
   563     if member op = ((map (Rule.thm_id)) rs) (Rule.thm_id r)
   564     then 
   565       let val ropt = Rewrite.rewrite_ thy ro eval_rls true (Rule.thm r) t;
   566       in 
   567         case ropt of
   568           SOME ta => [(r, ta)]
   569 	      | NONE => 
   570 	        ((*tracing ("### locate_rule:  rewrite " ^ Rule.thm_id r ^ " " ^ UnparseC.term t ^ " = NONE");*)
   571 	        []) end
   572     else ((*tracing ("### locate_rule:  " ^ Rule.thm_id r ^ " not mem rrls");*) [])
   573   | locate_rule _ _ _ _ _ _ = raise ERROR "locate_rule: doesnt match rev-sets in istate";
   574 
   575 fun next_rule thy eval_rls ro [rs] t =
   576     let val der = Derive.do_one thy eval_rls rs ro NONE t;
   577     in 
   578       case der of
   579 	      (_,r,_)::_ => SOME r
   580 	    | _ => NONE
   581     end
   582   | next_rule _ _ _ _ _ = raise ERROR ("next_rule: doesnt match rev-sets in istate");
   583 
   584 val pat0 = TermC.parse_patt \<^theory> "?r/?s+?u/?v :: real";
   585 val pat1 = TermC.parse_patt \<^theory> "?r/?s+?u    :: real";
   586 val pat2 = TermC.parse_patt \<^theory> "?r   +?u/?v :: real";
   587 val prepat = [([@{term True}], pat0),
   588 	      ([@{term True}], pat1),
   589 	      ([@{term True}], pat2)];
   590 (**)in(**)
   591 
   592 val add_fractions_p =
   593   Rule_Set.Rrls {id = "add_fractions_p", prepat=prepat,
   594     rew_ord = ("ord_make_polynomial", ord_make_polynomial false \<^theory>),
   595     erls = rational_erls,
   596     calc = [("PLUS", (\<^const_name>\<open>plus\<close>, (**)eval_binop "#add_")),
   597       ("TIMES", (\<^const_name>\<open>times\<close>, (**)eval_binop "#mult_")),
   598       ("DIVIDE", (\<^const_name>\<open>divide\<close>, Prog_Expr.eval_cancel "#divide_e")),
   599       ("POWER", (\<^const_name>\<open>powr\<close>, (**)eval_binop "#power_"))],
   600     errpatts = [],
   601     scr = Rule.Rfuns {init_state  = init_state \<^theory> Atools_erls ro,
   602       normal_form = add_fraction_p_ \<^theory>,
   603       locate_rule = locate_rule \<^theory> Atools_erls ro,
   604       next_rule   = next_rule \<^theory> Atools_erls ro,
   605       attach_form = attach_form}}
   606 (**)end(*local add_fractions_p *)
   607 \<close>
   608 
   609 subsection \<open>Cancelling and adding all occurrences in a term /////////////////////////////\<close>
   610 ML \<open>
   611 (*copying cancel_p_rls + add her caused error in interface.sml*)
   612 \<close>
   613 
   614 section \<open>Rulesets for general simplification\<close>
   615 ML \<open>
   616 (*.all powers over + distributed; atoms over * collected, other distributed
   617    contains absolute minimum of thms for context in norm_Rational .*)
   618 val powers = prep_rls'(
   619   Rule_Def.Repeat {id = "powers", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   620       erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   621       rules = [\<^rule_thm>\<open>realpow_multI\<close>,
   622 	       (*"(r * s)  \<up>  n = r  \<up>  n * s  \<up>  n"*)
   623 	       \<^rule_thm>\<open>realpow_pow\<close>,
   624 	       (*"(a  \<up>  b)  \<up>  c = a  \<up>  (b * c)"*)
   625 	       \<^rule_thm>\<open>realpow_oneI\<close>,
   626 	       (*"r  \<up>  1 = r"*)
   627 	       \<^rule_thm>\<open>realpow_minus_even\<close>,
   628 	       (*"n is_even ==> (- r)  \<up>  n = r  \<up>  n" ?-->discard_minus?*)
   629 	       \<^rule_thm>\<open>realpow_minus_odd\<close>,
   630 	       (*"Not (n is_even) ==> (- r)  \<up>  n = -1 * r  \<up>  n"*)
   631 	       
   632 	       (*----- collect atoms over * -----*)
   633 	       \<^rule_thm>\<open>realpow_two_atom\<close>,	
   634 	       (*"r is_atom ==> r * r = r  \<up>  2"*)
   635 	       \<^rule_thm>\<open>realpow_plus_1\<close>,		
   636 	       (*"r is_atom ==> r * r  \<up>  n = r  \<up>  (n + 1)"*)
   637 	       \<^rule_thm>\<open>realpow_addI_atom\<close>,
   638 	       (*"r is_atom ==> r  \<up>  n * r  \<up>  m = r  \<up>  (n + m)"*)
   639 
   640 	       (*----- distribute none-atoms -----*)
   641 	       \<^rule_thm>\<open>realpow_def_atom\<close>,
   642 	       (*"[| 1 < n; ~ (r is_atom) |]==>r  \<up>  n = r * r  \<up>  (n + -1)"*)
   643 	       \<^rule_thm>\<open>realpow_eq_oneI\<close>,
   644 	       (*"1  \<up>  n = 1"*)
   645 	       \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_")
   646 	       ],
   647       scr = Rule.Empty_Prog
   648       });
   649 (*.contains absolute minimum of thms for context in norm_Rational.*)
   650 val rat_mult_divide = prep_rls'(
   651   Rule_Def.Repeat {id = "rat_mult_divide", preconds = [], 
   652       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
   653       erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   654       rules = [\<^rule_thm>\<open>rat_mult\<close>,
   655 	       (*(1)"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
   656 	       \<^rule_thm>\<open>times_divide_eq_right\<close>,
   657 	       (*(2)"?a * (?c / ?d) = ?a * ?c / ?d" must be [2],
   658 	       otherwise inv.to a / b / c = ...*)
   659 	       \<^rule_thm>\<open>times_divide_eq_left\<close>,
   660 	       (*"?a / ?b * ?c = ?a * ?c / ?b" order weights x \<up> n too much
   661 		     and does not commute a / b * c  \<up>  2 !*)
   662 	       
   663 	       \<^rule_thm>\<open>divide_divide_eq_right\<close>,
   664 	       (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
   665 	       \<^rule_thm>\<open>divide_divide_eq_left\<close>,
   666 	       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
   667 	       \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e")
   668 	       ],
   669       scr = Rule.Empty_Prog
   670       });
   671 
   672 (*.contains absolute minimum of thms for context in norm_Rational.*)
   673 val reduce_0_1_2 = prep_rls'(
   674   Rule_Def.Repeat{id = "reduce_0_1_2", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   675       erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   676       rules = [(*\<^rule_thm>\<open>divide_1\<close>,
   677 		 "?x / 1 = ?x" unnecess.for normalform*)
   678 	       \<^rule_thm>\<open>mult_1_left\<close>,                 
   679 	       (*"1 * z = z"*)
   680 	       (*\<^rule_thm>\<open>real_mult_minus1\<close>,
   681 	       "-1 * z = - z"*)
   682 	       (*\<^rule_thm>\<open>real_minus_mult_cancel\<close>,
   683 	       "- ?x * - ?y = ?x * ?y"*)
   684 
   685 	       \<^rule_thm>\<open>mult_zero_left\<close>,        
   686 	       (*"0 * z = 0"*)
   687 	       \<^rule_thm>\<open>add_0_left\<close>,
   688 	       (*"0 + z = z"*)
   689 	       (*\<^rule_thm>\<open>right_minus\<close>,
   690 	       "?z + - ?z = 0"*)
   691 
   692 	       \<^rule_thm_sym>\<open>real_mult_2\<close>,	
   693 	       (*"z1 + z1 = 2 * z1"*)
   694 	       \<^rule_thm>\<open>real_mult_2_assoc\<close>,
   695 	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
   696 
   697 	       \<^rule_thm>\<open>division_ring_divide_zero\<close>
   698 	       (*"0 / ?x = 0"*)
   699 	       ], scr = Rule.Empty_Prog});
   700 
   701 (*erls for calculate_Rational; 
   702   make local with FIXX@ME result:term *term list WN0609???SKMG*)
   703 val norm_rat_erls = prep_rls'(
   704   Rule_Def.Repeat {id = "norm_rat_erls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   705       erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   706       rules = [\<^rule_eval>\<open>Prog_Expr.is_const\<close> (Prog_Expr.eval_const "#is_const_")
   707 	       ], scr = Rule.Empty_Prog});
   708 
   709 (* consists of rls containing the absolute minimum of thms *)
   710 (*040209: this version has been used by RL for his equations,
   711 which is now replaced by MGs version "norm_Rational" below *)
   712 val norm_Rational_min = prep_rls'(
   713   Rule_Def.Repeat {id = "norm_Rational_min", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   714       erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   715       rules = [(*sequence given by operator precedence*)
   716 	       Rule.Rls_ discard_minus,
   717 	       Rule.Rls_ powers,
   718 	       Rule.Rls_ rat_mult_divide,
   719 	       Rule.Rls_ expand,
   720 	       Rule.Rls_ reduce_0_1_2,
   721 	       Rule.Rls_ order_add_mult,
   722 	       Rule.Rls_ collect_numerals,
   723 	       Rule.Rls_ add_fractions_p,
   724 	       Rule.Rls_ cancel_p
   725 	       ],
   726       scr = Rule.Empty_Prog});
   727 
   728 val norm_Rational_parenthesized = prep_rls'(
   729   Rule_Set.Sequence {id = "norm_Rational_parenthesized", preconds = []:term list, 
   730        rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   731       erls = Atools_erls, srls = Rule_Set.Empty,
   732       calc = [], errpatts = [],
   733       rules = [Rule.Rls_  norm_Rational_min,
   734 	       Rule.Rls_ discard_parentheses
   735 	       ],
   736       scr = Rule.Empty_Prog});      
   737 
   738 (*WN030318???SK: simplifies all but cancel and common_nominator*)
   739 val simplify_rational = 
   740     Rule_Set.merge "simplify_rational" expand_binoms
   741     (Rule_Set.append_rules "divide" calculate_Rational
   742 		[\<^rule_thm>\<open>div_by_1\<close>,
   743 		 (*"?x / 1 = ?x"*)
   744 		 \<^rule_thm>\<open>rat_mult\<close>,
   745 		 (*(1)"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
   746 		 \<^rule_thm>\<open>times_divide_eq_right\<close>,
   747 		 (*(2)"?a * (?c / ?d) = ?a * ?c / ?d" must be [2],
   748 		 otherwise inv.to a / b / c = ...*)
   749 		 \<^rule_thm>\<open>times_divide_eq_left\<close>,
   750 		 (*"?a / ?b * ?c = ?a * ?c / ?b"*)
   751 		 \<^rule_thm>\<open>add_minus\<close>,
   752 		 (*"?a + ?b - ?b = ?a"*)
   753 		 \<^rule_thm>\<open>add_minus1\<close>,
   754 		 (*"?a - ?b + ?b = ?a"*)
   755 		 \<^rule_thm>\<open>divide_minus1\<close>
   756 		 (*"?x / -1 = - ?x"*)
   757 		 ]);
   758 \<close>
   759 ML \<open>
   760 val add_fractions_p_rls = prep_rls'(
   761   Rule_Def.Repeat {id = "add_fractions_p_rls", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
   762 	  erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   763 	  rules = [Rule.Rls_ add_fractions_p], 
   764 	  scr = Rule.Empty_Prog});
   765 
   766 (* "Rule_Def.Repeat" causes repeated application of cancel_p to one and the same term *)
   767 val cancel_p_rls = prep_rls'(
   768   Rule_Def.Repeat 
   769     {id = "cancel_p_rls", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
   770     erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   771     rules = [Rule.Rls_ cancel_p], 
   772 	  scr = Rule.Empty_Prog});
   773 
   774 (*. makes 'normal' fractions; 'is_polyexp' inhibits double fractions;
   775     used in initial part norm_Rational_mg, see example DA-M02-main.p.60.*)
   776 val rat_mult_poly = prep_rls'(
   777   Rule_Def.Repeat {id = "rat_mult_poly", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
   778 	  erls = Rule_Set.append_rules "Rule_Set.empty-is_polyexp" Rule_Set.empty
   779       [\<^rule_eval>\<open>is_polyexp\<close> (eval_is_polyexp "")],
   780 	  srls = Rule_Set.Empty, calc = [], errpatts = [],
   781 	  rules = 
   782 	    [\<^rule_thm>\<open>rat_mult_poly_l\<close>,
   783 	    (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
   784 	    \<^rule_thm>\<open>rat_mult_poly_r\<close>
   785 	    (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*) ], 
   786 	  scr = Rule.Empty_Prog});
   787 
   788 (*. makes 'normal' fractions; 'is_polyexp' inhibits double fractions;
   789     used in looping part norm_Rational_rls, see example DA-M02-main.p.60 
   790     .. WHERE THE LATTER DOES ALWAYS WORK, BECAUSE erls = Rule_Set.empty, 
   791     I.E. THE RESPECTIVE ASSUMPTION IS STORED AND Rule.Thm APPLIED; WN051028 
   792     ... WN0609???MG.*)
   793 val rat_mult_div_pow = prep_rls'(
   794   Rule_Def.Repeat {id = "rat_mult_div_pow", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   795     erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   796     rules = [\<^rule_thm>\<open>rat_mult\<close>,
   797       (*"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
   798       \<^rule_thm>\<open>rat_mult_poly_l\<close>,
   799       (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
   800       \<^rule_thm>\<open>rat_mult_poly_r\<close>,
   801       (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*)
   802       
   803       \<^rule_thm>\<open>real_divide_divide1_mg\<close>,
   804       (*"y ~= 0 ==> (u / v) / (y / z) = (u * z) / (y * v)"*)
   805       \<^rule_thm>\<open>divide_divide_eq_right\<close>,
   806       (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
   807       \<^rule_thm>\<open>divide_divide_eq_left\<close>,
   808       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
   809       \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
   810       
   811       \<^rule_thm>\<open>rat_power\<close>
   812       (*"(?a / ?b)  \<up>  ?n = ?a  \<up>  ?n / ?b  \<up>  ?n"*)
   813       ],
   814     scr = Rule.Empty_Prog});
   815 
   816 val rat_reduce_1 = prep_rls'(
   817   Rule_Def.Repeat {id = "rat_reduce_1", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
   818     erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [], 
   819     rules = 
   820       [\<^rule_thm>\<open>div_by_1\<close>,
   821       (*"?x / 1 = ?x"*)
   822       \<^rule_thm>\<open>mult_1_left\<close>           
   823       (*"1 * z = z"*)
   824       ],
   825     scr = Rule.Empty_Prog});
   826 
   827 (* looping part of norm_Rational *)
   828 val norm_Rational_rls = prep_rls' (
   829   Rule_Def.Repeat {id = "norm_Rational_rls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   830     erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   831     rules = [Rule.Rls_ add_fractions_p_rls,
   832       Rule.Rls_ rat_mult_div_pow,
   833       Rule.Rls_ make_rat_poly_with_parentheses,
   834       Rule.Rls_ cancel_p_rls,
   835       Rule.Rls_ rat_reduce_1
   836       ],
   837     scr = Rule.Empty_Prog});
   838 
   839 val norm_Rational = prep_rls' (
   840   Rule_Set.Sequence 
   841     {id = "norm_Rational", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
   842     erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   843     rules = [Rule.Rls_ discard_minus,
   844       Rule.Rls_ rat_mult_poly,             (* removes double fractions like a/b/c *)
   845       Rule.Rls_ make_rat_poly_with_parentheses,
   846       Rule.Rls_ cancel_p_rls,
   847       Rule.Rls_ norm_Rational_rls,         (* the main rls, looping (#) *)
   848       Rule.Rls_ discard_parentheses1       (* mult only *)
   849       ],
   850     scr = Rule.Empty_Prog});
   851 \<close>
   852 
   853 rule_set_knowledge
   854   calculate_Rational = calculate_Rational and
   855   calc_rat_erls = calc_rat_erls and
   856   rational_erls = rational_erls and
   857   cancel_p = cancel_p and
   858   add_fractions_p = add_fractions_p and
   859 
   860   add_fractions_p_rls = add_fractions_p_rls and
   861   powers_erls = powers_erls and
   862   powers = powers and
   863   rat_mult_divide = rat_mult_divide and
   864   reduce_0_1_2 = reduce_0_1_2 and
   865 
   866   rat_reduce_1 = rat_reduce_1 and
   867   norm_rat_erls = norm_rat_erls and
   868   norm_Rational = norm_Rational and
   869   norm_Rational_rls = norm_Rational_rls and
   870   norm_Rational_min = norm_Rational_min and
   871   norm_Rational_parenthesized = norm_Rational_parenthesized and
   872 
   873   rat_mult_poly = rat_mult_poly and
   874   rat_mult_div_pow = rat_mult_div_pow and
   875   cancel_p_rls = cancel_p_rls
   876 
   877 section \<open>A problem for simplification of rationals\<close>
   878 
   879 problem pbl_simp_rat : "rational/simplification" =
   880   \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
   881   Method: "simplification/of_rationals"
   882   CAS: "Simplify t_t"
   883   Given: "Term t_t"
   884   Where: "t_t is_ratpolyexp"
   885   Find: "normalform n_n"
   886 
   887 section \<open>A methods for simplification of rationals\<close>
   888 (*WN061025 this methods script is copied from (auto-generated) script
   889   of norm_Rational in order to ease repair on inform*)
   890 
   891 partial_function (tailrec) simplify :: "real \<Rightarrow> real"
   892   where
   893 "simplify term = (
   894   (Try (Rewrite_Set ''discard_minus'') #>
   895    Try (Rewrite_Set ''rat_mult_poly'') #>
   896    Try (Rewrite_Set ''make_rat_poly_with_parentheses'') #>
   897    Try (Rewrite_Set ''cancel_p_rls'') #>
   898    (Repeat (
   899      Try (Rewrite_Set ''add_fractions_p_rls'') #>
   900      Try (Rewrite_Set ''rat_mult_div_pow'') #>
   901      Try (Rewrite_Set ''make_rat_poly_with_parentheses'') #>
   902      Try (Rewrite_Set ''cancel_p_rls'') #>
   903      Try (Rewrite_Set ''rat_reduce_1''))) #>
   904    Try (Rewrite_Set ''discard_parentheses1''))
   905    term)"
   906 
   907 
   908 method met_simp_rat : "simplification/of_rationals" =
   909   \<open>{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, 
   910     prls = Rule_Set.append_rules "simplification_of_rationals_prls" Rule_Set.empty 
   911       [(*for preds in where_*) \<^rule_eval>\<open>is_ratpolyexp\<close> (eval_is_ratpolyexp "")],
   912     crls = Rule_Set.empty, errpats = [], nrls = norm_Rational_rls}\<close>
   913   Program: simplify.simps
   914   Given: "Term t_t"
   915   Where: "t_t is_ratpolyexp"
   916   Find: "normalform n_n"
   917 ML \<open>
   918 \<close> ML \<open>
   919 \<close>
   920 end