updated initial 8 thy's syntax without evaluation isac-update-Isa09-2
authorWalther Neuper <neuper@ist.tugraz.at>
Thu, 26 Aug 2010 18:15:30 +0200
branchisac-update-Isa09-2
changeset 37950525a28152a67
parent 37949 aaf528d3ebd5
child 37951 2c10fce11472
updated initial 8 thy's syntax without evaluation
src/Tools/isac/Knowledge/Equation.thy
src/Tools/isac/Knowledge/LinEq.thy
src/Tools/isac/Knowledge/Poly.thy
src/Tools/isac/Knowledge/PolyMinus.thy
src/Tools/isac/Knowledge/Rational.thy
src/Tools/isac/Knowledge/Root.thy
src/Tools/isac/Knowledge/RootEq.thy
src/Tools/isac/Knowledge/Simplify.thy
     1.1 --- a/src/Tools/isac/Knowledge/Equation.thy	Thu Aug 26 10:03:53 2010 +0200
     1.2 +++ b/src/Tools/isac/Knowledge/Equation.thy	Thu Aug 26 18:15:30 2010 +0200
     1.3 @@ -1,29 +1,100 @@
     1.4  (* equations and functions; functions NOT as lambda-terms
     1.5     author: Walther Neuper 2005, 2006
     1.6     (c) due to copyright terms
     1.7 -
     1.8 -remove_thy"Equation";
     1.9 -use_thy"Knowledge/Equation";
    1.10 -use_thy_only"Knowledge/Equation";
    1.11 -
    1.12 -remove_thy"Equation";
    1.13 -use_thy"Knowledge/Isac";
    1.14  *)
    1.15  
    1.16 -Equation = Atools +
    1.17 +theory Equation imports Atools begin
    1.18  
    1.19  consts
    1.20  
    1.21    (*descriptions in the related problems TODOshift here from Descriptions.thy*)
    1.22 -  substitution :: bool => una
    1.23 +  substitution      :: "bool => una"
    1.24  
    1.25    (*the CAS-commands*)
    1.26 -  solve     :: "[bool * 'a] => bool list" (* solve (x+1=2, x) *)
    1.27 -  solveTest :: "[bool * 'a] => bool list" (* for test collection *)
    1.28 +  solve             :: "[bool * 'a] => bool list" (* solve (x+1=2, x) *)
    1.29 +  solveTest         :: "[bool * 'a] => bool list" (* for test collection *)
    1.30    
    1.31    (*Script-names*)
    1.32 -  Function2Equality  :: "[bool, bool,       bool] \
    1.33 -					\=> bool"
    1.34 +  Function2Equality :: "[bool, bool,       bool] 
    1.35 +					 => bool"
    1.36                    ("((Script Function2Equality (_ _ =))// (_))" 9)
    1.37  
    1.38 +text {* defines equation and univariate-equation
    1.39 +        created by: rlang 
    1.40 +              date: 02.09
    1.41 +        changed by: rlang
    1.42 +        last change by: rlang
    1.43 +                  date: 02.11.29
    1.44 +        (c) by Richard Lang, 2003 *}
    1.45 +
    1.46 +ML {*
    1.47 +val univariate_equation_prls = 
    1.48 +    append_rls "univariate_equation_prls" e_rls 
    1.49 +	       [Calc ("Tools.matches",eval_matches "")];
    1.50 +ruleset' := 
    1.51 +overwritelthy thy (!ruleset',
    1.52 +		   [("univariate_equation_prls",
    1.53 +		     prep_rls univariate_equation_prls)]);
    1.54 +
    1.55 +
    1.56 +store_pbt 
    1.57 + (prep_pbt Equation.thy "pbl_equ" [] e_pblID
    1.58 + (["equation"],
    1.59 +  [("#Given" ,["equality e_","solveFor v_"]),
    1.60 +   ("#Where" ,["matches (?a = ?b) e_"]),
    1.61 +   ("#Find"  ,["solutions v_i_"])
    1.62 +  ],
    1.63 +  append_rls "equation_prls" e_rls 
    1.64 +	     [Calc ("Tools.matches",eval_matches "")],
    1.65 +  SOME "solve (e_::bool, v_)",
    1.66 +  []));
    1.67 +
    1.68 +store_pbt
    1.69 + (prep_pbt Equation.thy "pbl_equ_univ" [] e_pblID
    1.70 + (["univariate","equation"],
    1.71 +  [("#Given" ,["equality e_","solveFor v_"]),
    1.72 +   ("#Where" ,["matches (?a = ?b) e_"]),
    1.73 +   ("#Find"  ,["solutions v_i_"])
    1.74 +  ],
    1.75 +  univariate_equation_prls,SOME "solve (e_::bool, v_)",[]));
    1.76 +
    1.77 +
    1.78 +(*.function for handling the cas-input "solve (x+1=2, x)":
    1.79 +   make a model which is already in ptree-internal format.*)
    1.80 +(* val (h,argl) = strip_comb (str2term "solve (x+1=2, x)");
    1.81 +   val (h,argl) = strip_comb ((term_of o the o (parse thy)) 
    1.82 +				  "solveTest (x+1=2, x)");
    1.83 +   *)
    1.84 +fun argl2dtss [Const ("Pair", _) $ eq $ bdv] =
    1.85 +    [((term_of o the o (parse thy)) "equality", [eq]),
    1.86 +     ((term_of o the o (parse thy)) "solveFor", [bdv]),
    1.87 +     ((term_of o the o (parse thy)) "solutions", 
    1.88 +      [(term_of o the o (parse thy)) "L"])
    1.89 +     ]
    1.90 +  | argl2dtss _ = raise error "Equation.ML: wrong argument for argl2dtss";
    1.91 +
    1.92 +castab := 
    1.93 +overwritel (!castab, 
    1.94 +	    [((term_of o the o (parse thy)) "solveTest", 
    1.95 +	      (("Test.thy", ["univariate","equation","test"], ["no_met"]), 
    1.96 +	       argl2dtss)),
    1.97 +	     ((term_of o the o (parse thy)) "solve",  
    1.98 +	      (("Isac.thy", ["univariate","equation"], ["no_met"]), 
    1.99 +	       argl2dtss))
   1.100 +	     ]);
   1.101 +
   1.102 +
   1.103 +
   1.104 +store_met
   1.105 +    (prep_met Equation.thy "met_equ" [] e_metID
   1.106 +	      (["Equation"],
   1.107 +	       [],
   1.108 +	       {rew_ord'="tless_true", rls'=Erls, calc = [], 
   1.109 +		srls = e_rls, 
   1.110 +		prls=e_rls,
   1.111 +	     crls = Atools_erls, nrls = e_rls},
   1.112 +"empty_script"
   1.113 +));
   1.114 +*}
   1.115 +
   1.116  end
   1.117 \ No newline at end of file
     2.1 --- a/src/Tools/isac/Knowledge/LinEq.thy	Thu Aug 26 10:03:53 2010 +0200
     2.2 +++ b/src/Tools/isac/Knowledge/LinEq.thy	Thu Aug 26 18:15:30 2010 +0200
     2.3 @@ -7,44 +7,177 @@
     2.4               date: 02.10.20
     2.5  *)
     2.6  
     2.7 -(*
     2.8 - use"knowledge/LinEq.ML";
     2.9 - use"LinEq.ML";
    2.10 +theory LinEq imports Poly Equation begin
    2.11  
    2.12 - use"ROOT.ML";
    2.13 - cd"knowledge";
    2.14 -
    2.15 -*)
    2.16 -
    2.17 -LinEq = Poly + Equation +
    2.18 -
    2.19 -(*-------------------- consts------------------------------------------------*)
    2.20  consts
    2.21     Solve'_lineq'_equation
    2.22 -             :: "[bool,real, \
    2.23 -		  \ bool list] => bool list"
    2.24 -               ("((Script Solve'_lineq'_equation (_ _ =))// \
    2.25 -                 \ (_))" 9)
    2.26 +             :: "[bool,real, 
    2.27 +		   bool list] => bool list"
    2.28 +               ("((Script Solve'_lineq'_equation (_ _ =))// 
    2.29 +                  (_))" 9)
    2.30  
    2.31 -(*-------------------- rules -------------------------------------------------*)
    2.32 -rules
    2.33 +axioms
    2.34  (*-- normalize --*)
    2.35    (*WN0509 compare PolyEq.all_left "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"*)
    2.36 -  all_left
    2.37 -    "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)"
    2.38 -  makex1_x
    2.39 -    "a^^^1  = a"  
    2.40 -  real_assoc_1
    2.41 -   "a+(b+c) = a+b+c"
    2.42 -  real_assoc_2
    2.43 -   "a*(b*c) = a*b*c"
    2.44 +  all_left          "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)"
    2.45 +  makex1_x          "a^^^1  = a"  
    2.46 +  real_assoc_1      "a+(b+c) = a+b+c"
    2.47 +  real_assoc_2      "a*(b*c) = a*b*c"
    2.48  
    2.49  (*-- solve --*)
    2.50 -  lin_isolate_add1
    2.51 -   "(a + b*bdv = 0) = (b*bdv = (-1)*a)"
    2.52 -  lin_isolate_add2
    2.53 -   "(a +   bdv = 0) = (  bdv = (-1)*a)"
    2.54 -  lin_isolate_div
    2.55 -   "[|Not(b=0)|] ==> (b*bdv = c) = (bdv = c / b)"
    2.56 +  lin_isolate_add1  "(a + b*bdv = 0) = (b*bdv = (-1)*a)"
    2.57 +  lin_isolate_add2  "(a +   bdv = 0) = (  bdv = (-1)*a)"
    2.58 +  lin_isolate_div   "[|Not(b=0)|] ==> (b*bdv = c) = (bdv = c / b)"
    2.59 +
    2.60 +ML {*
    2.61 +val LinEq_prls = (*3.10.02:just the following order due to subterm evaluation*)
    2.62 +  append_rls "LinEq_prls" e_rls 
    2.63 +	     [Calc ("op =",eval_equal "#equal_"),
    2.64 +	      Calc ("Tools.matches",eval_matches ""),
    2.65 +	      Calc ("Tools.lhs"    ,eval_lhs ""),
    2.66 +	      Calc ("Tools.rhs"    ,eval_rhs ""),
    2.67 +	      Calc ("Poly.has'_degree'_in",eval_has_degree_in ""),
    2.68 + 	      Calc ("Poly.is'_polyrat'_in",eval_is_polyrat_in ""),
    2.69 +	      Calc ("Atools.occurs'_in",eval_occurs_in ""),    
    2.70 +	      Calc ("Atools.ident",eval_ident "#ident_"),
    2.71 +	      Thm ("not_true",num_str not_true),
    2.72 +	      Thm ("not_false",num_str not_false),
    2.73 +	      Thm ("and_true",num_str and_true),
    2.74 +	      Thm ("and_false",num_str and_false),
    2.75 +	      Thm ("or_true",num_str or_true),
    2.76 +	      Thm ("or_false",num_str or_false)
    2.77 +              ];
    2.78 +(* ----- erls ----- *)
    2.79 +val LinEq_crls = 
    2.80 +   append_rls "LinEq_crls" poly_crls
    2.81 +   [Thm  ("real_assoc_1",num_str real_assoc_1)
    2.82 +    (*		
    2.83 +     Don't use
    2.84 +     Calc ("HOL.divide", eval_cancel "#divide_"),
    2.85 +     Calc ("Atools.pow" ,eval_binop "#power_"),
    2.86 +     *)
    2.87 +    ];
    2.88 +
    2.89 +(* ----- crls ----- *)
    2.90 +val LinEq_erls = 
    2.91 +   append_rls "LinEq_erls" Poly_erls
    2.92 +   [Thm  ("real_assoc_1",num_str real_assoc_1)
    2.93 +    (*		
    2.94 +     Don't use
    2.95 +     Calc ("HOL.divide", eval_cancel "#divide_"),
    2.96 +     Calc ("Atools.pow" ,eval_binop "#power_"),
    2.97 +     *)
    2.98 +    ];
    2.99 +
   2.100 +ruleset' := overwritelthy thy (!ruleset',
   2.101 +			[("LinEq_erls",LinEq_erls)(*FIXXXME:del with rls.rls'*)
   2.102 +			 ]);
   2.103 +    
   2.104 +val LinPoly_simplify = prep_rls(
   2.105 +  Rls {id = "LinPoly_simplify", preconds = [], 
   2.106 +       rew_ord = ("termlessI",termlessI), 
   2.107 +       erls = LinEq_erls, 
   2.108 +       srls = Erls, 
   2.109 +       calc = [], 
   2.110 +       (*asm_thm = [],*)
   2.111 +       rules = [
   2.112 +		Thm  ("real_assoc_1",num_str real_assoc_1),
   2.113 +		Calc ("op +",eval_binop "#add_"),
   2.114 +		Calc ("op -",eval_binop "#sub_"),
   2.115 +		Calc ("op *",eval_binop "#mult_"),
   2.116 +		(*  Dont use  
   2.117 +		 Calc ("HOL.divide", eval_cancel "#divide_"),		
   2.118 +		 Calc ("Root.sqrt",eval_sqrt "#sqrt_"),
   2.119 +		 *)
   2.120 +		Calc ("Atools.pow" ,eval_binop "#power_")
   2.121 +		],
   2.122 +       scr = Script ((term_of o the o (parse thy)) "empty_script")
   2.123 +       }:rls);
   2.124 +ruleset' := overwritelthy thy (!ruleset',
   2.125 +			  [("LinPoly_simplify",LinPoly_simplify)]);
   2.126 +
   2.127 +(*isolate the bound variable in an linear equation; 'bdv' is a meta-constant*)
   2.128 +val LinEq_simplify = prep_rls(
   2.129 +Rls {id = "LinEq_simplify", preconds = [],
   2.130 +     rew_ord = ("e_rew_ord",e_rew_ord),
   2.131 +     erls = LinEq_erls,
   2.132 +     srls = Erls,
   2.133 +     calc = [],
   2.134 +     (*asm_thm = [("lin_isolate_div","")],*)
   2.135 +     rules = [
   2.136 +	      Thm("lin_isolate_add1",num_str lin_isolate_add1), 
   2.137 +	      (* a+bx=0 -> bx=-a *)
   2.138 +	      Thm("lin_isolate_add2",num_str lin_isolate_add2), 
   2.139 +	      (* a+ x=0 ->  x=-a *)
   2.140 +	      Thm("lin_isolate_div",num_str lin_isolate_div)    
   2.141 +	      (*   bx=c -> x=c/b *)  
   2.142 +	      ],
   2.143 +     scr = Script ((term_of o the o (parse thy)) "empty_script")
   2.144 +     }:rls);
   2.145 +ruleset' := overwritelthy thy (!ruleset',
   2.146 +			[("LinEq_simplify",LinEq_simplify)]);
   2.147 +
   2.148 +(*----------------------------- problem types --------------------------------*)
   2.149 +(* 
   2.150 +show_ptyps(); 
   2.151 +(get_pbt ["linear","univariate","equation"]);
   2.152 +*)
   2.153 +(* ---------linear----------- *)
   2.154 +store_pbt
   2.155 + (prep_pbt LinEq.thy "pbl_equ_univ_lin" [] e_pblID
   2.156 + (["linear","univariate","equation"],
   2.157 +  [("#Given" ,["equality e_","solveFor v_"]),
   2.158 +   ("#Where" ,["False", (*WN0509 just detected: this pbl can never be used?!?*)
   2.159 +               "Not( (lhs e_) is_polyrat_in v_)",
   2.160 +               "Not( (rhs e_) is_polyrat_in v_)",
   2.161 +               "((lhs e_) has_degree_in v_)=1",
   2.162 +	       "((rhs e_) has_degree_in v_)=1"]),
   2.163 +   ("#Find"  ,["solutions v_i_"]) 
   2.164 +  ],
   2.165 +  LinEq_prls, SOME "solve (e_::bool, v_)",
   2.166 +  [["LinEq","solve_lineq_equation"]]));
   2.167 +
   2.168 +(*-------------- methods------------------------------------------------------*)
   2.169 +store_met
   2.170 + (prep_met LinEq.thy "met_eqlin" [] e_metID
   2.171 + (["LinEq"],
   2.172 +   [],
   2.173 +   {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = e_rls, prls=e_rls,
   2.174 +    crls=LinEq_crls, nrls=norm_Poly
   2.175 +    (*, asm_rls=[],asm_thm=[]*)}, "empty_script"));
   2.176 +
   2.177 +(* ansprechen mit ["LinEq","solve_univar_equation"] *)
   2.178 +store_met
   2.179 +(prep_met LinEq.thy "met_eq_lin" [] e_metID
   2.180 + (["LinEq","solve_lineq_equation"],
   2.181 +   [("#Given" ,["equality e_","solveFor v_"]),
   2.182 +    ("#Where" ,["Not( (lhs e_) is_polyrat_in v_)",
   2.183 +                "( (lhs e_)  has_degree_in v_)=1"]),
   2.184 +    ("#Find"  ,["solutions v_i_"])
   2.185 +   ],
   2.186 +   {rew_ord'="termlessI",
   2.187 +    rls'=LinEq_erls,
   2.188 +    srls=e_rls,
   2.189 +    prls=LinEq_prls,
   2.190 +    calc=[],
   2.191 +    crls=LinEq_crls, nrls=norm_Poly(*,
   2.192 +    asm_rls=[],
   2.193 +    asm_thm=[("lin_isolate_div","")]*)},
   2.194 +    "Script Solve_lineq_equation (e_::bool) (v_::real) =                 " ^
   2.195 +    "(let e_ =((Try         (Rewrite     all_left            False)) @@  " ^ 
   2.196 +    "          (Try (Repeat (Rewrite     makex1_x           False))) @@  " ^ 
   2.197 +    "          (Try         (Rewrite_Set expand_binoms       False)) @@  " ^ 
   2.198 +    "          (Try (Repeat (Rewrite_Set_Inst [(bdv,v_::real)]           " ^
   2.199 +    "                                 make_ratpoly_in    False)))    @@  " ^
   2.200 +    "          (Try (Repeat (Rewrite_Set LinPoly_simplify      False)))) e_;" ^
   2.201 +    "     e_ = ((Try (Rewrite_Set_Inst [(bdv,v_::real)]                  " ^
   2.202 +    "                                          LinEq_simplify True)) @@  " ^
   2.203 +    "            (Repeat(Try (Rewrite_Set LinPoly_simplify     False)))) e_ " ^
   2.204 +    " in ((Or_to_List e_)::bool list))"
   2.205 + ));
   2.206 +"******* LinEq.ML end *******";
   2.207 +get_met ["LinEq","solve_lineq_equation"];
   2.208 +*}
   2.209 +
   2.210  end
   2.211  
     3.1 --- a/src/Tools/isac/Knowledge/Poly.thy	Thu Aug 26 10:03:53 2010 +0200
     3.2 +++ b/src/Tools/isac/Knowledge/Poly.thy	Thu Aug 26 18:15:30 2010 +0200
     3.3 @@ -7,41 +7,27 @@
     3.4     changed by: Richard Lang 020912
     3.5  *)
     3.6  
     3.7 -(*
     3.8 -   use_thy"Knowledge/Poly";
     3.9 -   use_thy"Poly";
    3.10 -   use_thy_only"Knowledge/Poly";
    3.11 +theory Poly imports Simplify begin
    3.12  
    3.13 -   remove_thy"Poly";
    3.14 -   use_thy"Knowledge/Isac";
    3.15 -
    3.16 -
    3.17 -   use"ROOT.ML";
    3.18 -   cd"IsacKnowledge";
    3.19 - *)
    3.20 -
    3.21 -Poly = Simplify + 
    3.22 -
    3.23 -(*-------------------- consts-----------------------------------------------*)
    3.24  consts
    3.25  
    3.26    is'_expanded'_in :: "[real, real] => bool" ("_ is'_expanded'_in _") 
    3.27 -  is'_poly'_in :: "[real, real] => bool" ("_ is'_poly'_in _")          (*RL DA *)
    3.28 -  has'_degree'_in :: "[real, real] => real" ("_ has'_degree'_in _")(*RL DA *)
    3.29 -  is'_polyrat'_in :: "[real, real] => bool" ("_ is'_polyrat'_in _")(*RL030626*)
    3.30 +  is'_poly'_in     :: "[real, real] => bool" ("_ is'_poly'_in _")   (*RL DA *)
    3.31 +  has'_degree'_in  :: "[real, real] => real" ("_ has'_degree'_in _")(*RL DA *)
    3.32 +  is'_polyrat'_in  :: "[real, real] => bool" ("_ is'_polyrat'_in _")(*RL030626*)
    3.33  
    3.34 - is'_multUnordered  :: "real => bool" ("_ is'_multUnordered") 
    3.35 - is'_addUnordered   :: "real => bool" ("_ is'_addUnordered") (*WN030618*)
    3.36 - is'_polyexp        :: "real => bool" ("_ is'_polyexp") 
    3.37 +  is'_multUnordered:: "real => bool" ("_ is'_multUnordered") 
    3.38 +  is'_addUnordered :: "real => bool" ("_ is'_addUnordered") (*WN030618*)
    3.39 +  is'_polyexp      :: "real => bool" ("_ is'_polyexp") 
    3.40  
    3.41    Expand'_binoms
    3.42 -             :: "['y, \
    3.43 -		  \ 'y] => 'y"
    3.44 -               ("((Script Expand'_binoms (_ =))// \
    3.45 -                 \ (_))" 9)
    3.46 +             :: "['y, 
    3.47 +		    'y] => 'y"
    3.48 +               ("((Script Expand'_binoms (_ =))// 
    3.49 +                    (_))" 9)
    3.50  
    3.51  (*-------------------- rules------------------------------------------------*)
    3.52 -rules (*.not contained in Isabelle2002,
    3.53 +axioms (*.not contained in Isabelle2002,
    3.54           stated as axioms, TODO: prove as theorems;
    3.55           theorem-IDs 'xxxI' with ^^^ instead of ^ in 'xxx' in Isabelle2002.*)
    3.56  
    3.57 @@ -54,8 +40,8 @@
    3.58    realpow_zeroI            "r ^^^ 0 = 1"
    3.59    realpow_eq_oneI         "1 ^^^ n = 1"
    3.60    realpow_multI           "(r * s) ^^^ n = r ^^^ n * s ^^^ n" 
    3.61 -  realpow_multI_poly      "[| r is_polyexp; s is_polyexp |] ==> \
    3.62 -			      \(r * s) ^^^ n = r ^^^ n * s ^^^ n" 
    3.63 +  realpow_multI_poly      "[| r is_polyexp; s is_polyexp |] ==>
    3.64 +			      (r * s) ^^^ n = r ^^^ n * s ^^^ n" 
    3.65    realpow_minus_oneI      "-1 ^^^ (2 * n) = 1"  
    3.66  
    3.67    realpow_twoI            "r ^^^ 2 = r * r"
    3.68 @@ -67,8 +53,8 @@
    3.69    realpow_plus_1_assoc_l2 "r ^^^ m * (r * s) = r ^^^ (1 + m) * s" 
    3.70    realpow_plus_1_assoc_r  "s * r * r ^^^ m = s * r ^^^ (1 + m)"
    3.71    realpow_plus_1_atom     "r is_atom ==> r * r ^^^ n = r ^^^ (1 + n)"
    3.72 -  realpow_def_atom        "[| Not (r is_atom); 1 < n |] \
    3.73 -			  \ ==> r ^^^ n = r * r ^^^ (n + -1)"
    3.74 +  realpow_def_atom        "[| Not (r is_atom); 1 < n |]
    3.75 +			   ==> r ^^^ n = r * r ^^^ (n + -1)"
    3.76    realpow_addI_atom       "r is_atom ==> r ^^^ n * r ^^^ m = r ^^^ (n + m)"
    3.77  
    3.78  
    3.79 @@ -77,35 +63,38 @@
    3.80  
    3.81  
    3.82  (* RL 020914 *)
    3.83 -  real_pp_binom_times        "(a + b)*(c + d) = a*c + a*d + b*c + b*d"
    3.84 -  real_pm_binom_times        "(a + b)*(c - d) = a*c - a*d + b*c - b*d"
    3.85 -  real_mp_binom_times        "(a - b)*(c + d) = a*c + a*d - b*c - b*d"
    3.86 -  real_mm_binom_times        "(a - b)*(c - d) = a*c - a*d - b*c + b*d"
    3.87 -  real_plus_binom_pow3       "(a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3"
    3.88 -  real_plus_binom_pow3_poly  "[| a is_polyexp; b is_polyexp |] ==> \
    3.89 -			      \(a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3"
    3.90 -  real_minus_binom_pow3      "(a - b)^^^3 = a^^^3 - 3*a^^^2*b + 3*a*b^^^2 - b^^^3"
    3.91 -  real_minus_binom_pow3_p    "(a + -1 * b)^^^3 = a^^^3 + -3*a^^^2*b + 3*a*b^^^2 + -1*b^^^3"
    3.92 -(* real_plus_binom_pow        "[| n is_const;  3 < n |] ==>  \
    3.93 -			      \(a + b)^^^n = (a + b) * (a + b)^^^(n - 1)" *)
    3.94 -  real_plus_binom_pow4       "(a + b)^^^4 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)*(a + b)"
    3.95 -  real_plus_binom_pow4_poly  "[| a is_polyexp; b is_polyexp |] ==> \
    3.96 -			      \(a + b)^^^4 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)*(a + b)"
    3.97 -  real_plus_binom_pow5       "(a + b)^^^5 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)*(a^^^2 + 2*a*b + b^^^2)"
    3.98 -
    3.99 -  real_plus_binom_pow5_poly  "[| a is_polyexp; b is_polyexp |] ==> \
   3.100 -			      \(a + b)^^^5 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)*(a^^^2 + 2*a*b + b^^^2)"
   3.101 -
   3.102 -  real_diff_plus             "a - b = a + -b" (*17.3.03: do_NOT_use*)
   3.103 -  real_diff_minus            "a - b = a + -1 * b"
   3.104 -  real_plus_binom_times      "(a + b)*(a + b) = a^^^2 + 2*a*b + b^^^2"
   3.105 -  real_minus_binom_times     "(a - b)*(a - b) = a^^^2 - 2*a*b + b^^^2"
   3.106 +  real_pp_binom_times     "(a + b)*(c + d) = a*c + a*d + b*c + b*d"
   3.107 +  real_pm_binom_times     "(a + b)*(c - d) = a*c - a*d + b*c - b*d"
   3.108 +  real_mp_binom_times     "(a - b)*(c + d) = a*c + a*d - b*c - b*d"
   3.109 +  real_mm_binom_times     "(a - b)*(c - d) = a*c - a*d - b*c + b*d"
   3.110 +  real_plus_binom_pow3    "(a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3"
   3.111 +  real_plus_binom_pow3_poly "[| a is_polyexp; b is_polyexp |] ==> 
   3.112 +			    (a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3"
   3.113 +  real_minus_binom_pow3   "(a - b)^^^3 = a^^^3 - 3*a^^^2*b + 3*a*b^^^2 - b^^^3"
   3.114 +  real_minus_binom_pow3_p "(a + -1 * b)^^^3 = a^^^3 + -3*a^^^2*b + 3*a*b^^^2 +
   3.115 +                           -1*b^^^3"
   3.116 +(* real_plus_binom_pow        "[| n is_const;  3 < n |] ==>
   3.117 +			       (a + b)^^^n = (a + b) * (a + b)^^^(n - 1)" *)
   3.118 +  real_plus_binom_pow4    "(a + b)^^^4 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)
   3.119 +                           *(a + b)"
   3.120 +  real_plus_binom_pow4_poly "[| a is_polyexp; b is_polyexp |] ==> 
   3.121 +			   (a + b)^^^4 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)
   3.122 +                           *(a + b)"
   3.123 +  real_plus_binom_pow5    "(a + b)^^^5 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3)
   3.124 +                           *(a^^^2 + 2*a*b + b^^^2)"
   3.125 +  real_plus_binom_pow5_poly "[| a is_polyexp; b is_polyexp |] ==> 
   3.126 +			        (a + b)^^^5 = (a^^^3 + 3*a^^^2*b + 3*a*b^^^2 
   3.127 +                                + b^^^3)*(a^^^2 + 2*a*b + b^^^2)"
   3.128 +  real_diff_plus          "a - b = a + -b" (*17.3.03: do_NOT_use*)
   3.129 +  real_diff_minus         "a - b = a + -1 * b"
   3.130 +  real_plus_binom_times   "(a + b)*(a + b) = a^^^2 + 2*a*b + b^^^2"
   3.131 +  real_minus_binom_times  "(a - b)*(a - b) = a^^^2 - 2*a*b + b^^^2"
   3.132    (*WN071229 changed for Schaerding -----vvv*)
   3.133 -  (*real_plus_binom_pow2       "(a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"*)
   3.134 -  real_plus_binom_pow2       "(a + b)^^^2 = (a + b) * (a + b)"
   3.135 +  (*real_plus_binom_pow2  "(a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"*)
   3.136 +  real_plus_binom_pow2    "(a + b)^^^2 = (a + b) * (a + b)"
   3.137    (*WN071229 changed for Schaerding -----^^^*)
   3.138 -  real_plus_binom_pow2_poly   "[| a is_polyexp; b is_polyexp |] ==> \
   3.139 -			      \(a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"
   3.140 +  real_plus_binom_pow2_poly "[| a is_polyexp; b is_polyexp |] ==>
   3.141 +			       (a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"
   3.142    real_minus_binom_pow2      "(a - b)^^^2 = a^^^2 - 2*a*b + b^^^2"
   3.143    real_minus_binom_pow2_p    "(a - b)^^^2 = a^^^2 + -2*a*b + b^^^2"
   3.144    real_plus_minus_binom1     "(a + b)*(a - b) = a^^^2 - b^^^2"
   3.145 @@ -117,24 +106,24 @@
   3.146    real_plus_binom_times1     "(a +  1*b)*(a + -1*b) = a^^^2 + -1*b^^^2"
   3.147    real_plus_binom_times2     "(a + -1*b)*(a +  1*b) = a^^^2 + -1*b^^^2"
   3.148  
   3.149 -  real_num_collect           "[| l is_const; m is_const |] ==> \
   3.150 -					\l * n + m * n = (l + m) * n"
   3.151 +  real_num_collect           "[| l is_const; m is_const |] ==>
   3.152 +			      l * n + m * n = (l + m) * n"
   3.153  (* FIXME.MG.0401: replace 'real_num_collect_assoc' 
   3.154  	by 'real_num_collect_assoc_l' ... are equal, introduced by MG ! *)
   3.155 -  real_num_collect_assoc     "[| l is_const; m is_const |] ==>  \
   3.156 -					\l * n + (m * n + k) = (l + m) * n + k"
   3.157 -  real_num_collect_assoc_l     "[| l is_const; m is_const |] ==>  \
   3.158 -					\l * n + (m * n + k) = (l + m)
   3.159 -					* n + k"
   3.160 -  real_num_collect_assoc_r     "[| l is_const; m is_const |] ==>  \
   3.161 -					\(k + m * n) + l * n = k + (l + m) * n"
   3.162 +  real_num_collect_assoc     "[| l is_const; m is_const |] ==> 
   3.163 +			      l * n + (m * n + k) = (l + m) * n + k"
   3.164 +  real_num_collect_assoc_l   "[| l is_const; m is_const |] ==>
   3.165 +			      l * n + (m * n + k) = (l + m)
   3.166 +				* n + k"
   3.167 +  real_num_collect_assoc_r   "[| l is_const; m is_const |] ==>
   3.168 +			      (k + m * n) + l * n = k + (l + m) * n"
   3.169    real_one_collect           "m is_const ==> n + m * n = (1 + m) * n"
   3.170  (* FIXME.MG.0401: replace 'real_one_collect_assoc' 
   3.171  	by 'real_one_collect_assoc_l' ... are equal, introduced by MG ! *)
   3.172    real_one_collect_assoc     "m is_const ==> n + (m * n + k) = (1 + m)* n + k"
   3.173  
   3.174    real_one_collect_assoc_l   "m is_const ==> n + (m * n + k) = (1 + m) * n + k"
   3.175 -  real_one_collect_assoc_r   "m is_const ==>(k + n) +  m * n = k + (1 + m) * n"
   3.176 +  real_one_collect_assoc_r   "m is_const ==> (k + n) +  m * n = k + (1 + m) * n"
   3.177  
   3.178  (* FIXME.MG.0401: replace 'real_mult_2_assoc' 
   3.179  	by 'real_mult_2_assoc_l' ... are equal, introduced by MG ! *)
   3.180 @@ -144,4 +133,1484 @@
   3.181  
   3.182    real_add_mult_distrib_poly "w is_polyexp ==> (z1 + z2) * w = z1 * w + z2 * w"
   3.183    real_add_mult_distrib2_poly "w is_polyexp ==> w * (z1 + z2) = w * z1 + w * z2"
   3.184 +
   3.185 +text {* remark on 'polynomials'
   3.186 +        WN020919
   3.187 +*** there are 5 kinds of expanded normalforms ***
   3.188 +
   3.189 +[1] 'complete polynomial' (Komplettes Polynom), univariate
   3.190 +   a_0 + a_1.x^1 +...+ a_n.x^n   not (a_n = 0)
   3.191 +	        not (a_n = 0), some a_i may be zero (DON'T disappear),
   3.192 +                variables in monomials lexicographically ordered and complete,
   3.193 +                x written as 1*x^1, ...
   3.194 +[2] 'polynomial' (Polynom), univariate and multivariate
   3.195 +   a_0 + a_1.x +...+ a_n.x^n   not (a_n = 0)
   3.196 +   a_0 + a_1.x_1.x_2^n_12...x_m^n_1m +...+  a_n.x_1^n.x_2^n_n2...x_m^n_nm
   3.197 +	        not (a_n = 0), some a_i may be zero (ie. monomials disappear),
   3.198 +                exponents and coefficients equal 1 are not (WN060904.TODO in cancel_p_)shown,
   3.199 +                and variables in monomials are lexicographically ordered  
   3.200 +   examples: [1]: "1 + (-10) * x ^^^ 1 + 25 * x ^^^ 2"
   3.201 +	     [1]: "11 + 0 * x ^^^ 1 + 1 * x ^^^ 2"
   3.202 +	     [2]: "x + (-50) * x ^^^ 3"
   3.203 +	     [2]: "(-1) * x * y ^^^ 2 + 7 * x ^^^ 3"
   3.204 +
   3.205 +[3] 'expanded_term' (Ausmultiplizierter Term):
   3.206 +   pull out unary minus to binary minus, 
   3.207 +   as frequently exercised in schools; other conditions for [2] hold however
   3.208 +   examples: "a ^^^ 2 - 2 * a * b + b ^^^ 2"
   3.209 +	     "4 * x ^^^ 2 - 9 * y ^^^ 2"
   3.210 +[4] 'polynomial_in' (Polynom in): 
   3.211 +   polynomial in 1 variable with arbitrary coefficients
   3.212 +   examples: "2 * x + (-50) * x ^^^ 3"                     (poly in x)
   3.213 +	     "(u + v) + (2 * u ^^^ 2) * a + (-u) * a ^^^ 2 (poly in a)
   3.214 +[5] 'expanded_in' (Ausmultiplizierter Termin in): 
   3.215 +   analoguous to [3] with binary minus like [3]
   3.216 +   examples: "2 * x - 50 * x ^^^ 3"                     (expanded in x)
   3.217 +	     "(u + v) + (2 * u ^^^ 2) * a - u * a ^^^ 2 (expanded in a)
   3.218 +*}
   3.219 +
   3.220 +ML {*
   3.221 +(* is_polyrat_in becomes true, if no bdv is in the denominator of a fraction*)
   3.222 +fun is_polyrat_in t v = 
   3.223 +    let fun coeff_in c v = member op = (vars c) v;
   3.224 +   	fun finddivide (_ $ _ $ _ $ _) v = raise error("is_polyrat_in:")
   3.225 +	    (* at the moment there is no term like this, but ....*)
   3.226 +	  | finddivide (t as (Const ("HOL.divide",_) $ _ $ b)) v = 
   3.227 +            not(coeff_in b v)
   3.228 +	  | finddivide (_ $ t1 $ t2) v = 
   3.229 +            (finddivide t1 v) orelse (finddivide t2 v)
   3.230 +	  | finddivide (_ $ t1) v = (finddivide t1 v)
   3.231 +	  | finddivide _ _ = false;
   3.232 +     in finddivide t v end;
   3.233 +    
   3.234 +fun eval_is_polyrat_in _ _(p as (Const ("Poly.is'_polyrat'_in",_) $ t $ v)) _  =
   3.235 +    if is_polyrat_in t v 
   3.236 +    then SOME ((term2str p) ^ " = True",
   3.237 +	        Trueprop $ (mk_equality (p, HOLogic.true_const)))
   3.238 +    else SOME ((term2str p) ^ " = True",
   3.239 +	        Trueprop $ (mk_equality (p, HOLogic.false_const)))
   3.240 +  | eval_is_polyrat_in _ _ _ _ = ((*writeln"### no matches";*) NONE);
   3.241 +
   3.242 +local
   3.243 +    (*.a 'c is coefficient of v' if v does NOT occur in c.*)
   3.244 +    fun coeff_in c v = not (member op = (vars c) v);
   3.245 +    (* FIXME.WN100826 shift this into test--------------
   3.246 +     val v = (term_of o the o (parse thy)) "x";
   3.247 +     val t = (term_of o the o (parse thy)) "1";
   3.248 +     coeff_in t v;
   3.249 +     (*val it = true : bool*)
   3.250 +     val t = (term_of o the o (parse thy)) "a*b+c";
   3.251 +     coeff_in t v;
   3.252 +     (*val it = true : bool*)
   3.253 +     val t = (term_of o the o (parse thy)) "a*x+c";
   3.254 +     coeff_in t v;
   3.255 +     (*val it = false : bool*)
   3.256 +    ----------------------------------------------------*)
   3.257 +    (*. a 'monomial t in variable v' is a term t with
   3.258 +      either (1) v NOT existent in t, or (2) v contained in t,
   3.259 +      if (1) then degree 0
   3.260 +      if (2) then v is a factor on the very right, ev. with exponent.*)
   3.261 +    fun factor_right_deg (*case 2*)
   3.262 +    	    (t as Const ("op *",_) $ t1 $ 
   3.263 +    	       (Const ("Atools.pow",_) $ vv $ Free (d,_))) v =
   3.264 +    	if ((vv = v) andalso (coeff_in t1 v)) then SOME (int_of_str' d) else NONE
   3.265 +      | factor_right_deg (t as Const ("Atools.pow",_) $ vv $ Free (d,_)) v =
   3.266 +    	if (vv = v) then SOME (int_of_str' d) else NONE
   3.267 +      | factor_right_deg (t as Const ("op *",_) $ t1 $ vv) v = 
   3.268 +    	if ((vv = v) andalso (coeff_in t1 v))then SOME 1 else NONE
   3.269 +      | factor_right_deg vv v =
   3.270 +    	if (vv = v) then SOME 1 else NONE;    
   3.271 +    fun mono_deg_in m v =
   3.272 +    	if coeff_in m v then (*case 1*) SOME 0
   3.273 +    	else factor_right_deg m v;
   3.274 +    (* FIXME.WN100826 shift this into test-----------------------------
   3.275 +     val v = (term_of o the o (parse thy)) "x";
   3.276 +     val t = (term_of o the o (parse thy)) "(a*b+c)*x^^^7";
   3.277 +     mono_deg_in t v;
   3.278 +     (*val it = SOME 7*)
   3.279 +     val t = (term_of o the o (parse thy)) "x^^^7";
   3.280 +     mono_deg_in t v;
   3.281 +     (*val it = SOME 7*)
   3.282 +     val t = (term_of o the o (parse thy)) "(a*b+c)*x";
   3.283 +     mono_deg_in t v;
   3.284 +     (*val it = SOME 1*)
   3.285 +     val t = (term_of o the o (parse thy)) "(a*b+x)*x";
   3.286 +     mono_deg_in t v;
   3.287 +     (*val it = NONE*)
   3.288 +     val t = (term_of o the o (parse thy)) "x";
   3.289 +     mono_deg_in t v;
   3.290 +     (*val it = SOME 1*)
   3.291 +     val t = (term_of o the o (parse thy)) "(a*b+c)";
   3.292 +     mono_deg_in t v;
   3.293 +     (*val it = SOME 0*)
   3.294 +     val t = (term_of o the o (parse thy)) "ab - (a*b)*x";
   3.295 +     mono_deg_in t v;
   3.296 +     (*val it = NONE*)
   3.297 +    ------------------------------------------------------------------*)
   3.298 +    fun expand_deg_in t v =
   3.299 +    	let fun edi ~1 ~1 (Const ("op +",_) $ t1 $ t2) =
   3.300 +    		(case mono_deg_in t2 v of (* $ is left associative*)
   3.301 +    		     SOME d' => edi d' d' t1
   3.302 +		   | NONE => NONE)
   3.303 +    	      | edi ~1 ~1 (Const ("op -",_) $ t1 $ t2) =
   3.304 +    		(case mono_deg_in t2 v of
   3.305 +    		     SOME d' => edi d' d' t1
   3.306 +		   | NONE => NONE)
   3.307 +    	      | edi d dmax (Const ("op -",_) $ t1 $ t2) =
   3.308 +    		(case mono_deg_in t2 v of
   3.309 +		(*RL  orelse ((d=0) andalso (d'=0)) need to handle 3+4-...4 +x*)
   3.310 +    		     SOME d' => if ((d > d') orelse ((d=0) andalso (d'=0))) 
   3.311 +                     then edi d' dmax t1 else NONE
   3.312 +		   | NONE => NONE)
   3.313 +    	      | edi d dmax (Const ("op +",_) $ t1 $ t2) =
   3.314 +    		(case mono_deg_in t2 v of
   3.315 +		(*RL  orelse ((d=0) andalso (d'=0)) need to handle 3+4-...4 +x*)
   3.316 +    		     SOME d' => if ((d > d') orelse ((d=0) andalso (d'=0))) 
   3.317 +                     then edi d' dmax t1 else NONE
   3.318 +		   | NONE => NONE)
   3.319 +    	      | edi ~1 ~1 t = (case mono_deg_in t v of
   3.320 +    		     d as SOME _ => d
   3.321 +		   | NONE => NONE)
   3.322 +    	      | edi d dmax t = (*basecase last*)
   3.323 +    		(case mono_deg_in t v of
   3.324 +    		     SOME d' => if ((d > d') orelse ((d=0) andalso (d'=0)))  
   3.325 +                     then SOME dmax else NONE
   3.326 +		   | NONE => NONE)
   3.327 +    	in edi ~1 ~1 t end;
   3.328 +    (* FIXME.WN100826 shift this into test-----------------------------
   3.329 +     val v = (term_of o the o (parse thy)) "x";
   3.330 +     val t = (term_of o the o (parse thy)) "a+b";
   3.331 +     expand_deg_in t v;
   3.332 +     (*val it = SOME 0*)   
   3.333 +     val t = (term_of o the o (parse thy)) "(a+b)*x";
   3.334 +     expand_deg_in t v;
   3.335 +     (*SOME 1*)   
   3.336 +     val t = (term_of o the o (parse thy)) "a*b - (a+b)*x";
   3.337 +     expand_deg_in t v;
   3.338 +     (*SOME 1*)   
   3.339 +     val t = (term_of o the o (parse thy)) "a*b + (a-b)*x";
   3.340 +     expand_deg_in t v;
   3.341 +     (*SOME 1*)   
   3.342 +     val t = (term_of o the o (parse thy)) "a*b + (a+b)*x + x^^^2";
   3.343 +     expand_deg_in t v;
   3.344 +    -------------------------------------------------------------------*)   
   3.345 +    fun poly_deg_in t v =
   3.346 +    	let fun edi ~1 ~1 (Const ("op +",_) $ t1 $ t2) =
   3.347 +    		(case mono_deg_in t2 v of (* $ is left associative*)
   3.348 +    		     SOME d' => edi d' d' t1
   3.349 +		   | NONE => NONE)
   3.350 +    	      | edi d dmax (Const ("op +",_) $ t1 $ t2) =
   3.351 +    		(case mono_deg_in t2 v of
   3.352 + 		(*RL  orelse ((d=0) andalso (d'=0)) need to handle 3+4-...4 +x*)
   3.353 +   		     SOME d' => if ((d > d') orelse ((d=0) andalso (d'=0))) 
   3.354 +                                then edi d' dmax t1 else NONE
   3.355 +		   | NONE => NONE)
   3.356 +    	      | edi ~1 ~1 t = (case mono_deg_in t v of
   3.357 +    		     d as SOME _ => d
   3.358 +		   | NONE => NONE)
   3.359 +    	      | edi d dmax t = (*basecase last*)
   3.360 +    		(case mono_deg_in t v of
   3.361 +    		     SOME d' => if ((d > d') orelse ((d=0) andalso (d'=0))) 
   3.362 +                     then SOME dmax else NONE
   3.363 +		   | NONE => NONE)
   3.364 +    	in edi ~1 ~1 t end;
   3.365 +in
   3.366 +
   3.367 +fun is_expanded_in t v =
   3.368 +    case expand_deg_in t v of SOME _ => true | NONE => false;
   3.369 +fun is_poly_in t v =
   3.370 +    case poly_deg_in t v of SOME _ => true | NONE => false;
   3.371 +fun has_degree_in t v =
   3.372 +    case expand_deg_in t v of SOME d => d | NONE => ~1;
   3.373 +end;
   3.374 +(* FIXME.WN100826 shift this into test-----------------------------
   3.375 + val v = (term_of o the o (parse thy)) "x";
   3.376 + val t = (term_of o the o (parse thy)) "a*b - (a+b)*x + x^^^2";
   3.377 + has_degree_in t v;
   3.378 + (*val it = 2*)
   3.379 + val t = (term_of o the o (parse thy)) "-8 - 2*x + x^^^2";
   3.380 + has_degree_in t v;
   3.381 + (*val it = 2*)
   3.382 + val t = (term_of o the o (parse thy)) "6 + 13*x + 6*x^^^2";
   3.383 + has_degree_in t v;
   3.384 + (*val it = 2*)
   3.385 +-------------------------------------------------------------------*)
   3.386 +
   3.387 +(*("is_expanded_in", ("Poly.is'_expanded'_in", eval_is_expanded_in ""))*)
   3.388 +fun eval_is_expanded_in _ _ 
   3.389 +       (p as (Const ("Poly.is'_expanded'_in",_) $ t $ v)) _ =
   3.390 +    if is_expanded_in t v
   3.391 +    then SOME ((term2str p) ^ " = True",
   3.392 +	        Trueprop $ (mk_equality (p, HOLogic.true_const)))
   3.393 +    else SOME ((term2str p) ^ " = True",
   3.394 +	        Trueprop $ (mk_equality (p, HOLogic.false_const)))
   3.395 +  | eval_is_expanded_in _ _ _ _ = NONE;
   3.396 +(*
   3.397 + val t = (term_of o the o (parse thy)) "(-8 - 2*x + x^^^2) is_expanded_in x";
   3.398 + val SOME (id, t') = eval_is_expanded_in 0 0 t 0;
   3.399 + (*val id = "Poly.is'_expanded'_in (-8 - 2 * x + x ^^^ 2) x = True"*)
   3.400 + term2str t';
   3.401 + (*val it = "Poly.is'_expanded'_in (-8 - 2 * x + x ^^^ 2) x = True"*)
   3.402 +*)
   3.403 +
   3.404 +(*("is_poly_in", ("Poly.is'_poly'_in", eval_is_poly_in ""))*)
   3.405 +fun eval_is_poly_in _ _ 
   3.406 +       (p as (Const ("Poly.is'_poly'_in",_) $ t $ v)) _ =
   3.407 +    if is_poly_in t v
   3.408 +    then SOME ((term2str p) ^ " = True",
   3.409 +	        Trueprop $ (mk_equality (p, HOLogic.true_const)))
   3.410 +    else SOME ((term2str p) ^ " = True",
   3.411 +	        Trueprop $ (mk_equality (p, HOLogic.false_const)))
   3.412 +  | eval_is_poly_in _ _ _ _ = NONE;
   3.413 +(*
   3.414 + val t = (term_of o the o (parse thy)) "(8 + 2*x + x^^^2) is_poly_in x";
   3.415 + val SOME (id, t') = eval_is_poly_in 0 0 t 0;
   3.416 + (*val id = "Poly.is'_poly'_in (8 + 2 * x + x ^^^ 2) x = True"*)
   3.417 + term2str t';
   3.418 + (*val it = "Poly.is'_poly'_in (8 + 2 * x + x ^^^ 2) x = True"*)
   3.419 +*)
   3.420 +
   3.421 +(*("has_degree_in", ("Poly.has'_degree'_in", eval_has_degree_in ""))*)
   3.422 +fun eval_has_degree_in _ _ 
   3.423 +	     (p as (Const ("Poly.has'_degree'_in",_) $ t $ v)) _ =
   3.424 +    let val d = has_degree_in t v
   3.425 +	val d' = term_of_num HOLogic.realT d
   3.426 +    in SOME ((term2str p) ^ " = " ^ (string_of_int d),
   3.427 +	      Trueprop $ (mk_equality (p, d')))
   3.428 +    end
   3.429 +  | eval_has_degree_in _ _ _ _ = NONE;
   3.430 +(*
   3.431 +> val t = (term_of o the o (parse thy)) "(-8 - 2*x + x^^^2) has_degree_in x";
   3.432 +> val SOME (id, t') = eval_has_degree_in 0 0 t 0;
   3.433 +val id = "Poly.has'_degree'_in (-8 - 2 * x + x ^^^ 2) x = 2" : string
   3.434 +> term2str t';
   3.435 +val it = "Poly.has'_degree'_in (-8 - 2 * x + x ^^^ 2) x = 2" : string
   3.436 +*)
   3.437 +
   3.438 +(*.for evaluation of conditions in rewrite rules.*)
   3.439 +val Poly_erls = 
   3.440 +    append_rls "Poly_erls" Atools_erls
   3.441 +               [ Calc ("op =",eval_equal "#equal_"),
   3.442 +		 Thm  ("real_unari_minus",num_str real_unari_minus),
   3.443 +                 Calc ("op +",eval_binop "#add_"),
   3.444 +		 Calc ("op -",eval_binop "#sub_"),
   3.445 +		 Calc ("op *",eval_binop "#mult_"),
   3.446 +		 Calc ("Atools.pow" ,eval_binop "#power_")
   3.447 +		 ];
   3.448 +
   3.449 +val poly_crls = 
   3.450 +    append_rls "poly_crls" Atools_crls
   3.451 +               [ Calc ("op =",eval_equal "#equal_"),
   3.452 +		 Thm  ("real_unari_minus",num_str real_unari_minus),
   3.453 +                 Calc ("op +",eval_binop "#add_"),
   3.454 +		 Calc ("op -",eval_binop "#sub_"),
   3.455 +		 Calc ("op *",eval_binop "#mult_"),
   3.456 +		 Calc ("Atools.pow" ,eval_binop "#power_")
   3.457 +		 ];
   3.458 +
   3.459 +local (*. for make_polynomial .*)
   3.460 +
   3.461 +open Term;  (* for type order = EQUAL | LESS | GREATER *)
   3.462 +
   3.463 +fun pr_ord EQUAL = "EQUAL"
   3.464 +  | pr_ord LESS  = "LESS"
   3.465 +  | pr_ord GREATER = "GREATER";
   3.466 +
   3.467 +fun dest_hd' (Const (a, T)) =                          (* ~ term.ML *)
   3.468 +  (case a of
   3.469 +     "Atools.pow" => ((("|||||||||||||", 0), T), 0)    (*WN greatest string*)
   3.470 +   | _ => (((a, 0), T), 0))
   3.471 +  | dest_hd' (Free (a, T)) = (((a, 0), T), 1)
   3.472 +  | dest_hd' (Var v) = (v, 2)
   3.473 +  | dest_hd' (Bound i) = ((("", i), dummyT), 3)
   3.474 +  | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4);
   3.475 +
   3.476 +fun get_order_pow (t $ (Free(order,_))) = (* RL FIXXXME:geht zufaellig?WN*)
   3.477 +    	(case int_of_str (order) of
   3.478 +	             SOME d => d
   3.479 +		   | NONE   => 0)
   3.480 +  | get_order_pow _ = 0;
   3.481 +
   3.482 +fun size_of_term' (Const(str,_) $ t) =
   3.483 +  if "Atools.pow"= str then 1000 + size_of_term' t else 1+size_of_term' t(*WN*)
   3.484 +  | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
   3.485 +  | size_of_term' (f$t) = size_of_term' f  +  size_of_term' t
   3.486 +  | size_of_term' _ = 1;
   3.487 +
   3.488 +fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
   3.489 +      (case term_ord' pr thy (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
   3.490 +  | term_ord' pr thy (t, u) =
   3.491 +      (if pr then 
   3.492 +	 let
   3.493 +	   val (f, ts) = strip_comb t and (g, us) = strip_comb u;
   3.494 +	   val _=writeln("t= f@ts= \""^
   3.495 +	      ((Syntax.string_of_term (thy2ctxt thy)) f)^"\" @ \"["^
   3.496 +	      (commas(map(Syntax.string_of_term (thy2ctxt thy))ts))^"]\"");
   3.497 +	   val _=writeln("u= g@us= \""^
   3.498 +	      ((Syntax.string_of_term (thy2ctxt thy)) g)^"\" @ \"["^
   3.499 +	      (commas(map(Syntax.string_of_term (thy2ctxt thy))us))^"]\"");
   3.500 +	   val _=writeln("size_of_term(t,u)= ("^
   3.501 +	      (string_of_int(size_of_term' t))^", "^
   3.502 +	      (string_of_int(size_of_term' u))^")");
   3.503 +	   val _=writeln("hd_ord(f,g)      = "^((pr_ord o hd_ord)(f,g)));
   3.504 +	   val _=writeln("terms_ord(ts,us) = "^
   3.505 +			   ((pr_ord o terms_ord str false)(ts,us)));
   3.506 +	   val _=writeln("-------");
   3.507 +	 in () end
   3.508 +       else ();
   3.509 +	 case int_ord (size_of_term' t, size_of_term' u) of
   3.510 +	   EQUAL =>
   3.511 +	     let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
   3.512 +	       (case hd_ord (f, g) of EQUAL => (terms_ord str pr) (ts, us) 
   3.513 +	     | ord => ord)
   3.514 +	     end
   3.515 +	 | ord => ord)
   3.516 +and hd_ord (f, g) =                                        (* ~ term.ML *)
   3.517 +  prod_ord (prod_ord indexname_ord typ_ord) int_ord (dest_hd' f, dest_hd' g)
   3.518 +and terms_ord str pr (ts, us) = 
   3.519 +    list_ord (term_ord' pr (assoc_thy "Isac.thy"))(ts, us);
   3.520 +in
   3.521 +
   3.522 +fun ord_make_polynomial (pr:bool) thy (_:subst) tu = 
   3.523 +    (term_ord' pr thy(***) tu = LESS );
   3.524 +
   3.525 +end;(*local*)
   3.526 +
   3.527 +
   3.528 +rew_ord' := overwritel (!rew_ord',
   3.529 +[("termlessI", termlessI),
   3.530 + ("ord_make_polynomial", ord_make_polynomial false thy)
   3.531 + ]);
   3.532 +
   3.533 +
   3.534 +val expand =
   3.535 +  Rls{id = "expand", preconds = [], 
   3.536 +      rew_ord = ("dummy_ord", dummy_ord),
   3.537 +      erls = e_rls,srls = Erls,
   3.538 +      calc = [],
   3.539 +      (*asm_thm = [],*)
   3.540 +      rules = [Thm ("real_add_mult_distrib" ,num_str real_add_mult_distrib),
   3.541 +	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   3.542 +	       Thm ("real_add_mult_distrib2",num_str real_add_mult_distrib2)
   3.543 +	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   3.544 +	       ], scr = EmptyScr}:rls;
   3.545 +
   3.546 +(*----------------- Begin: rulesets for make_polynomial_ -----------------
   3.547 +  'rlsIDs' redefined by MG as 'rlsIDs_' 
   3.548 +                                    ^^^*)
   3.549 +
   3.550 +val discard_minus_ = 
   3.551 +  Rls{id = "discard_minus_", preconds = [], 
   3.552 +      rew_ord = ("dummy_ord", dummy_ord),
   3.553 +      erls = e_rls,srls = Erls,
   3.554 +      calc = [],
   3.555 +      (*asm_thm = [],*)
   3.556 +      rules = [Thm ("real_diff_minus",num_str real_diff_minus),
   3.557 +	       (*"a - b = a + -1 * b"*)
   3.558 +	       Thm ("sym_real_mult_minus1",num_str (real_mult_minus1 RS sym))
   3.559 +	       (*- ?z = "-1 * ?z"*)
   3.560 +	       ], scr = EmptyScr}:rls;
   3.561 +val expand_poly_ = 
   3.562 +  Rls{id = "expand_poly_", preconds = [], 
   3.563 +      rew_ord = ("dummy_ord", dummy_ord),
   3.564 +      erls = e_rls,srls = Erls,
   3.565 +      calc = [],
   3.566 +      (*asm_thm = [],*)
   3.567 +      rules = [Thm ("real_plus_binom_pow4",num_str real_plus_binom_pow4),
   3.568 +	       (*"(a + b)^^^4 = ... "*)
   3.569 +	       Thm ("real_plus_binom_pow5",num_str real_plus_binom_pow5),
   3.570 +	       (*"(a + b)^^^5 = ... "*)
   3.571 +	       Thm ("real_plus_binom_pow3",num_str real_plus_binom_pow3),
   3.572 +	       (*"(a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3" *)
   3.573 +
   3.574 +	       (*WN071229 changed/removed for Schaerding -----vvv*)
   3.575 +	       (*Thm ("real_plus_binom_pow2",num_str real_plus_binom_pow2),*)
   3.576 +	       (*"(a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"*)
   3.577 +	       Thm ("real_plus_binom_pow2",num_str real_plus_binom_pow2),
   3.578 +	       (*"(a + b)^^^2 = (a + b) * (a + b)"*)
   3.579 +	       (*Thm ("real_plus_minus_binom1_p_p",
   3.580 +		    num_str real_plus_minus_binom1_p_p),*)
   3.581 +	       (*"(a + b)*(a + -1 * b) = a^^^2 + -1*b^^^2"*)
   3.582 +	       (*Thm ("real_plus_minus_binom2_p_p",
   3.583 +		    num_str real_plus_minus_binom2_p_p),*)
   3.584 +	       (*"(a + -1 * b)*(a + b) = a^^^2 + -1*b^^^2"*)
   3.585 +	       (*WN071229 changed/removed for Schaerding -----^^^*)
   3.586 +	      
   3.587 +	       Thm ("real_add_mult_distrib" ,num_str real_add_mult_distrib),
   3.588 +	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   3.589 +	       Thm ("real_add_mult_distrib2",num_str real_add_mult_distrib2),
   3.590 +	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   3.591 +	       
   3.592 +	       Thm ("realpow_multI", num_str realpow_multI),
   3.593 +	       (*"(r * s) ^^^ n = r ^^^ n * s ^^^ n"*)
   3.594 +	       Thm ("realpow_pow",num_str realpow_pow)
   3.595 +	       (*"(a ^^^ b) ^^^ c = a ^^^ (b * c)"*)
   3.596 +	       ], scr = EmptyScr}:rls;
   3.597 +
   3.598 +(*.the expression contains + - * ^ only ?
   3.599 +   this is weaker than 'is_polynomial' !.*)
   3.600 +fun is_polyexp (Free _) = true
   3.601 +  | is_polyexp (Const ("op +",_) $ Free _ $ Free _) = true
   3.602 +  | is_polyexp (Const ("op -",_) $ Free _ $ Free _) = true
   3.603 +  | is_polyexp (Const ("op *",_) $ Free _ $ Free _) = true
   3.604 +  | is_polyexp (Const ("Atools.pow",_) $ Free _ $ Free _) = true
   3.605 +  | is_polyexp (Const ("op +",_) $ t1 $ t2) = 
   3.606 +               ((is_polyexp t1) andalso (is_polyexp t2))
   3.607 +  | is_polyexp (Const ("op -",_) $ t1 $ t2) = 
   3.608 +               ((is_polyexp t1) andalso (is_polyexp t2))
   3.609 +  | is_polyexp (Const ("op *",_) $ t1 $ t2) = 
   3.610 +               ((is_polyexp t1) andalso (is_polyexp t2))
   3.611 +  | is_polyexp (Const ("Atools.pow",_) $ t1 $ t2) = 
   3.612 +               ((is_polyexp t1) andalso (is_polyexp t2))
   3.613 +  | is_polyexp _ = false;
   3.614 +
   3.615 +(*("is_polyexp", ("Poly.is'_polyexp", eval_is_polyexp ""))*)
   3.616 +fun eval_is_polyexp (thmid:string) _ 
   3.617 +		       (t as (Const("Poly.is'_polyexp", _) $ arg)) thy = 
   3.618 +    if is_polyexp arg
   3.619 +    then SOME (mk_thmid thmid "" 
   3.620 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
   3.621 +	       Trueprop $ (mk_equality (t, HOLogic.true_const)))
   3.622 +    else SOME (mk_thmid thmid "" 
   3.623 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
   3.624 +	       Trueprop $ (mk_equality (t, HOLogic.false_const)))
   3.625 +  | eval_is_polyexp _ _ _ _ = NONE; 
   3.626 +
   3.627 +val expand_poly_rat_ = 
   3.628 +  Rls{id = "expand_poly_rat_", preconds = [], 
   3.629 +      rew_ord = ("dummy_ord", dummy_ord),
   3.630 +      erls =  append_rls "e_rls-is_polyexp" e_rls
   3.631 +	        [Calc ("Poly.is'_polyexp", eval_is_polyexp "")
   3.632 +		 ],
   3.633 +      srls = Erls,
   3.634 +      calc = [],
   3.635 +      (*asm_thm = [],*)
   3.636 +      rules = 
   3.637 +        [Thm ("real_plus_binom_pow4_poly", num_str real_plus_binom_pow4_poly),
   3.638 +	     (*"[| a is_polyexp; b is_polyexp |] ==> (a + b)^^^4 = ... "*)
   3.639 +	 Thm ("real_plus_binom_pow5_poly", num_str real_plus_binom_pow5_poly),
   3.640 +	     (*"[| a is_polyexp; b is_polyexp |] ==> (a + b)^^^5 = ... "*)
   3.641 +	 Thm ("real_plus_binom_pow2_poly",num_str real_plus_binom_pow2_poly),
   3.642 +	     (*"[| a is_polyexp; b is_polyexp |] ==>
   3.643 +		            (a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"*)
   3.644 +	 Thm ("real_plus_binom_pow3_poly",num_str real_plus_binom_pow3_poly),
   3.645 +	     (*"[| a is_polyexp; b is_polyexp |] ==> 
   3.646 +			(a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3" *)
   3.647 +	 Thm ("real_plus_minus_binom1_p_p",num_str real_plus_minus_binom1_p_p),
   3.648 +	     (*"(a + b)*(a + -1 * b) = a^^^2 + -1*b^^^2"*)
   3.649 +	 Thm ("real_plus_minus_binom2_p_p",num_str real_plus_minus_binom2_p_p),
   3.650 +	     (*"(a + -1 * b)*(a + b) = a^^^2 + -1*b^^^2"*)
   3.651 +	      
   3.652 +	 Thm ("real_add_mult_distrib_poly" ,num_str real_add_mult_distrib_poly),
   3.653 +	       (*"w is_polyexp ==> (z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   3.654 +	 Thm("real_add_mult_distrib2_poly",num_str real_add_mult_distrib2_poly),
   3.655 +	     (*"w is_polyexp ==> w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   3.656 +	       
   3.657 +	 Thm ("realpow_multI_poly", num_str realpow_multI_poly),
   3.658 +	     (*"[| r is_polyexp; s is_polyexp |] ==> 
   3.659 +		            (r * s) ^^^ n = r ^^^ n * s ^^^ n"*)
   3.660 +	  Thm ("realpow_pow",num_str realpow_pow)
   3.661 +	      (*"(a ^^^ b) ^^^ c = a ^^^ (b * c)"*)
   3.662 +	 ], scr = EmptyScr}:rls;
   3.663 +
   3.664 +val simplify_power_ = 
   3.665 +  Rls{id = "simplify_power_", preconds = [], 
   3.666 +      rew_ord = ("dummy_ord", dummy_ord),
   3.667 +      erls = e_rls, srls = Erls,
   3.668 +      calc = [],
   3.669 +      (*asm_thm = [],*)
   3.670 +      rules = [(*MG: Reihenfolge der folgenden 2 Thm muss so bleiben, wegen
   3.671 +		a*(a*a) --> a*a^^^2 und nicht a*(a*a) --> a^^^2*a *)
   3.672 +	       Thm ("sym_realpow_twoI",num_str (realpow_twoI RS sym)),	
   3.673 +	       (*"r * r = r ^^^ 2"*)
   3.674 +	       Thm ("realpow_twoI_assoc_l",num_str realpow_twoI_assoc_l),
   3.675 +	       (*"r * (r * s) = r ^^^ 2 * s"*)
   3.676 +
   3.677 +	       Thm ("realpow_plus_1",num_str realpow_plus_1),		
   3.678 +	       (*"r * r ^^^ n = r ^^^ (n + 1)"*)
   3.679 +	       Thm ("realpow_plus_1_assoc_l", num_str realpow_plus_1_assoc_l),
   3.680 +	       (*"r * (r ^^^ m * s) = r ^^^ (1 + m) * s"*)
   3.681 +	       (*MG 9.7.03: neues Thm wegen a*(a*(a*b)) --> a^^^2*(a*b) *)
   3.682 +	       Thm ("realpow_plus_1_assoc_l2", num_str realpow_plus_1_assoc_l2),
   3.683 +	       (*"r ^^^ m * (r * s) = r ^^^ (1 + m) * s"*)
   3.684 +
   3.685 +	       Thm ("sym_realpow_addI",num_str (realpow_addI RS sym)),
   3.686 +	       (*"r ^^^ n * r ^^^ m = r ^^^ (n + m)"*)
   3.687 +	       Thm ("realpow_addI_assoc_l", num_str realpow_addI_assoc_l),
   3.688 +	       (*"r ^^^ n * (r ^^^ m * s) = r ^^^ (n + m) * s"*)
   3.689 +	       
   3.690 +	       (* ist in expand_poly - wird hier aber auch gebraucht, wegen: 
   3.691 +		  "r * r = r ^^^ 2" wenn r=a^^^b*)
   3.692 +	       Thm ("realpow_pow",num_str realpow_pow)
   3.693 +	       (*"(a ^^^ b) ^^^ c = a ^^^ (b * c)"*)
   3.694 +	       ], scr = EmptyScr}:rls;
   3.695 +
   3.696 +val calc_add_mult_pow_ = 
   3.697 +  Rls{id = "calc_add_mult_pow_", preconds = [], 
   3.698 +      rew_ord = ("dummy_ord", dummy_ord),
   3.699 +      erls = Atools_erls(*erls3.4.03*),srls = Erls,
   3.700 +      calc = [("PLUS"  , ("op +", eval_binop "#add_")), 
   3.701 +	      ("TIMES" , ("op *", eval_binop "#mult_")),
   3.702 +	      ("POWER", ("Atools.pow", eval_binop "#power_"))
   3.703 +	      ],
   3.704 +      (*asm_thm = [],*)
   3.705 +      rules = [Calc ("op +", eval_binop "#add_"),
   3.706 +	       Calc ("op *", eval_binop "#mult_"),
   3.707 +	       Calc ("Atools.pow", eval_binop "#power_")
   3.708 +	       ], scr = EmptyScr}:rls;
   3.709 +
   3.710 +val reduce_012_mult_ = 
   3.711 +  Rls{id = "reduce_012_mult_", preconds = [], 
   3.712 +      rew_ord = ("dummy_ord", dummy_ord),
   3.713 +      erls = e_rls,srls = Erls,
   3.714 +      calc = [],
   3.715 +      (*asm_thm = [],*)
   3.716 +      rules = [(* MG: folgende Thm müssen hier stehen bleiben: *)
   3.717 +               Thm ("real_mult_1_right",num_str real_mult_1_right),
   3.718 +	       (*"z * 1 = z"*) (*wegen "a * b * b^^^(-1) + a"*) 
   3.719 +	       Thm ("realpow_zeroI",num_str realpow_zeroI),
   3.720 +	       (*"r ^^^ 0 = 1"*) (*wegen "a*a^^^(-1)*c + b + c"*)
   3.721 +	       Thm ("realpow_oneI",num_str realpow_oneI),
   3.722 +	       (*"r ^^^ 1 = r"*)
   3.723 +	       Thm ("realpow_eq_oneI",num_str realpow_eq_oneI)
   3.724 +	       (*"1 ^^^ n = 1"*)
   3.725 +	       ], scr = EmptyScr}:rls;
   3.726 +
   3.727 +val collect_numerals_ = 
   3.728 +  Rls{id = "collect_numerals_", preconds = [], 
   3.729 +      rew_ord = ("dummy_ord", dummy_ord),
   3.730 +      erls = Atools_erls, srls = Erls,
   3.731 +      calc = [("PLUS"  , ("op +", eval_binop "#add_"))
   3.732 +	      ],
   3.733 +      rules = 
   3.734 +        [Thm ("real_num_collect",num_str real_num_collect), 
   3.735 +	     (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
   3.736 +	 Thm ("real_num_collect_assoc_r",num_str real_num_collect_assoc_r),
   3.737 +	     (*"[| l is_const; m is_const |] ==>  \
   3.738 +					\(k + m * n) + l * n = k + (l + m)*n"*)
   3.739 +	 Thm ("real_one_collect",num_str real_one_collect),	
   3.740 +	     (*"m is_const ==> n + m * n = (1 + m) * n"*)
   3.741 +	 Thm ("real_one_collect_assoc_r",num_str real_one_collect_assoc_r), 
   3.742 +	     (*"m is_const ==> (k + n) + m * n = k + (m + 1) * n"*)
   3.743 +
   3.744 +         Calc ("op +", eval_binop "#add_"),
   3.745 +
   3.746 +	 (*MG: Reihenfolge der folgenden 2 Thm muss so bleiben, wegen
   3.747 +		     (a+a)+a --> a + 2*a --> 3*a and not (a+a)+a --> 2*a + a *)
   3.748 +         Thm ("real_mult_2_assoc_r",num_str real_mult_2_assoc_r),
   3.749 +	     (*"(k + z1) + z1 = k + 2 * z1"*)
   3.750 +	 Thm ("sym_real_mult_2",num_str (real_mult_2 RS sym))
   3.751 +	     (*"z1 + z1 = 2 * z1"*)
   3.752 +	], scr = EmptyScr}:rls;
   3.753 +
   3.754 +val reduce_012_ = 
   3.755 +  Rls{id = "reduce_012_", preconds = [], 
   3.756 +      rew_ord = ("dummy_ord", dummy_ord),
   3.757 +      erls = e_rls,srls = Erls,
   3.758 +      calc = [],
   3.759 +      (*asm_thm = [],*)
   3.760 +      rules = [Thm ("real_mult_1",num_str real_mult_1),                 
   3.761 +	       (*"1 * z = z"*)
   3.762 +	       Thm ("real_mult_0",num_str real_mult_0),        
   3.763 +	       (*"0 * z = 0"*)
   3.764 +	       Thm ("real_mult_0_right",num_str real_mult_0_right),        
   3.765 +	       (*"z * 0 = 0"*)
   3.766 +	       Thm ("real_add_zero_left",num_str real_add_zero_left),
   3.767 +	       (*"0 + z = z"*)
   3.768 +	       Thm ("real_add_zero_right",num_str real_add_zero_right),
   3.769 +	       (*"z + 0 = z"*) (*wegen a+b-b --> a+(1-1)*b --> a+0 --> a*)
   3.770 +
   3.771 +	       (*Thm ("realpow_oneI",num_str realpow_oneI)*)
   3.772 +	       (*"?r ^^^ 1 = ?r"*)
   3.773 +	       Thm ("real_0_divide",num_str real_0_divide)(*WN060914*)
   3.774 +	       (*"0 / ?x = 0"*)
   3.775 +	       ], scr = EmptyScr}:rls;
   3.776 +
   3.777 +(*ein Hilfs-'ruleset' (benutzt das leere 'ruleset')*)
   3.778 +val discard_parentheses_ = 
   3.779 +    append_rls "discard_parentheses_" e_rls 
   3.780 +	       [Thm ("sym_real_mult_assoc", num_str (real_mult_assoc RS sym))
   3.781 +		(*"?z1.1 * (?z2.1 * ?z3.1) = ?z1.1 * ?z2.1 * ?z3.1"*)
   3.782 +		(*Thm ("sym_real_add_assoc",num_str (real_add_assoc RS sym))*)
   3.783 +		(*"?z1.1 + (?z2.1 + ?z3.1) = ?z1.1 + ?z2.1 + ?z3.1"*)
   3.784 +		 ];
   3.785 +
   3.786 +(*----------------- End: rulesets for make_polynomial_ -----------------*)
   3.787 +
   3.788 +(*MG.0401 ev. for use in rls with ordered rewriting ?
   3.789 +val collect_numerals_left = 
   3.790 +  Rls{id = "collect_numerals", preconds = [], 
   3.791 +      rew_ord = ("dummy_ord", dummy_ord),
   3.792 +      erls = Atools_erls(*erls3.4.03*),srls = Erls,
   3.793 +      calc = [("PLUS"  , ("op +", eval_binop "#add_")), 
   3.794 +	      ("TIMES" , ("op *", eval_binop "#mult_")),
   3.795 +	      ("POWER", ("Atools.pow", eval_binop "#power_"))
   3.796 +	      ],
   3.797 +      (*asm_thm = [],*)
   3.798 +      rules = [Thm ("real_num_collect",num_str real_num_collect), 
   3.799 +	       (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
   3.800 +	       Thm ("real_num_collect_assoc",num_str real_num_collect_assoc),
   3.801 +	       (*"[| l is_const; m is_const |] ==>  
   3.802 +				l * n + (m * n + k) =  (l + m) * n + k"*)
   3.803 +	       Thm ("real_one_collect",num_str real_one_collect),	
   3.804 +	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
   3.805 +	       Thm ("real_one_collect_assoc",num_str real_one_collect_assoc), 
   3.806 +	       (*"m is_const ==> n + (m * n + k) = (1 + m) * n + k"*)
   3.807 +	       
   3.808 +	       Calc ("op +", eval_binop "#add_"),
   3.809 +
   3.810 +	       (*MG am 2.5.03: 2 Theoreme aus reduce_012 hierher verschoben*)
   3.811 +	       Thm ("sym_real_mult_2",num_str (real_mult_2 RS sym)),	
   3.812 +	       (*"z1 + z1 = 2 * z1"*)
   3.813 +	       Thm ("real_mult_2_assoc",num_str real_mult_2_assoc)
   3.814 +	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
   3.815 +	       ], scr = EmptyScr}:rls;*)
   3.816 +
   3.817 +val expand_poly = 
   3.818 +  Rls{id = "expand_poly", preconds = [], 
   3.819 +      rew_ord = ("dummy_ord", dummy_ord),
   3.820 +      erls = e_rls,srls = Erls,
   3.821 +      calc = [],
   3.822 +      (*asm_thm = [],*)
   3.823 +      rules = [Thm ("real_add_mult_distrib" ,num_str real_add_mult_distrib),
   3.824 +	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   3.825 +	       Thm ("real_add_mult_distrib2",num_str real_add_mult_distrib2),
   3.826 +	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   3.827 +	       (*Thm ("real_add_mult_distrib1",num_str real_add_mult_distrib1),
   3.828 +		....... 18.3.03 undefined???*)
   3.829 +
   3.830 +	       Thm ("real_plus_binom_pow2",num_str real_plus_binom_pow2),
   3.831 +	       (*"(a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"*)
   3.832 +	       Thm ("real_minus_binom_pow2_p",num_str real_minus_binom_pow2_p),
   3.833 +	       (*"(a - b)^^^2 = a^^^2 + -2*a*b + b^^^2"*)
   3.834 +	       Thm ("real_plus_minus_binom1_p",
   3.835 +		    num_str real_plus_minus_binom1_p),
   3.836 +	       (*"(a + b)*(a - b) = a^^^2 + -1*b^^^2"*)
   3.837 +	       Thm ("real_plus_minus_binom2_p",
   3.838 +		    num_str real_plus_minus_binom2_p),
   3.839 +	       (*"(a - b)*(a + b) = a^^^2 + -1*b^^^2"*)
   3.840 +
   3.841 +	       Thm ("real_minus_minus",num_str real_minus_minus),
   3.842 +	       (*"- (- ?z) = ?z"*)
   3.843 +	       Thm ("real_diff_minus",num_str real_diff_minus),
   3.844 +	       (*"a - b = a + -1 * b"*)
   3.845 +	       Thm ("sym_real_mult_minus1",num_str (real_mult_minus1 RS sym))
   3.846 +	       (*- ?z = "-1 * ?z"*)
   3.847 +
   3.848 +	       (*Thm ("",num_str ),
   3.849 +	       Thm ("",num_str ),
   3.850 +	       Thm ("",num_str ),*)
   3.851 +	       (*Thm ("real_minus_add_distrib",
   3.852 +		      num_str real_minus_add_distrib),*)
   3.853 +	       (*"- (?x + ?y) = - ?x + - ?y"*)
   3.854 +	       (*Thm ("real_diff_plus",num_str real_diff_plus)*)
   3.855 +	       (*"a - b = a + -b"*)
   3.856 +	       ], scr = EmptyScr}:rls;
   3.857 +
   3.858 +val simplify_power = 
   3.859 +  Rls{id = "simplify_power", preconds = [], 
   3.860 +      rew_ord = ("dummy_ord", dummy_ord),
   3.861 +      erls = e_rls, srls = Erls,
   3.862 +      calc = [],
   3.863 +      (*asm_thm = [],*)
   3.864 +      rules = [Thm ("realpow_multI", num_str realpow_multI),
   3.865 +	       (*"(r * s) ^^^ n = r ^^^ n * s ^^^ n"*)
   3.866 +	       
   3.867 +	       Thm ("sym_realpow_twoI",num_str (realpow_twoI RS sym)),	
   3.868 +	       (*"r1 * r1 = r1 ^^^ 2"*)
   3.869 +	       Thm ("realpow_plus_1",num_str realpow_plus_1),		
   3.870 +	       (*"r * r ^^^ n = r ^^^ (n + 1)"*)
   3.871 +	       Thm ("realpow_pow",num_str realpow_pow),
   3.872 +	       (*"(a ^^^ b) ^^^ c = a ^^^ (b * c)"*)
   3.873 +	       Thm ("sym_realpow_addI",num_str (realpow_addI RS sym)),
   3.874 +	       (*"r ^^^ n * r ^^^ m = r ^^^ (n + m)"*)
   3.875 +	       Thm ("realpow_oneI",num_str realpow_oneI),
   3.876 +	       (*"r ^^^ 1 = r"*)
   3.877 +	       Thm ("realpow_eq_oneI",num_str realpow_eq_oneI)
   3.878 +	       (*"1 ^^^ n = 1"*)
   3.879 +	       ], scr = EmptyScr}:rls;
   3.880 +(*MG.0401: termorders for multivariate polys dropped due to principal problems:
   3.881 +  (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*)
   3.882 +val order_add_mult = 
   3.883 +  Rls{id = "order_add_mult", preconds = [], 
   3.884 +      rew_ord = ("ord_make_polynomial",ord_make_polynomial false Poly.thy),
   3.885 +      erls = e_rls,srls = Erls,
   3.886 +      calc = [],
   3.887 +      (*asm_thm = [],*)
   3.888 +      rules = [Thm ("real_mult_commute",num_str real_mult_commute),
   3.889 +	       (* z * w = w * z *)
   3.890 +	       Thm ("real_mult_left_commute",num_str real_mult_left_commute),
   3.891 +	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
   3.892 +	       Thm ("real_mult_assoc",num_str real_mult_assoc),		
   3.893 +	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
   3.894 +	       Thm ("real_add_commute",num_str real_add_commute),	
   3.895 +	       (*z + w = w + z*)
   3.896 +	       Thm ("real_add_left_commute",num_str real_add_left_commute),
   3.897 +	       (*x + (y + z) = y + (x + z)*)
   3.898 +	       Thm ("real_add_assoc",num_str real_add_assoc)	               
   3.899 +	       (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
   3.900 +	       ], scr = EmptyScr}:rls;
   3.901 +(*MG.0401: termorders for multivariate polys dropped due to principal problems:
   3.902 +  (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*)
   3.903 +val order_mult = 
   3.904 +  Rls{id = "order_mult", preconds = [], 
   3.905 +      rew_ord = ("ord_make_polynomial",ord_make_polynomial false Poly.thy),
   3.906 +      erls = e_rls,srls = Erls,
   3.907 +      calc = [],
   3.908 +      (*asm_thm = [],*)
   3.909 +      rules = [Thm ("real_mult_commute",num_str real_mult_commute),
   3.910 +	       (* z * w = w * z *)
   3.911 +	       Thm ("real_mult_left_commute",num_str real_mult_left_commute),
   3.912 +	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
   3.913 +	       Thm ("real_mult_assoc",num_str real_mult_assoc)	
   3.914 +	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
   3.915 +	       ], scr = EmptyScr}:rls;
   3.916 +val collect_numerals = 
   3.917 +  Rls{id = "collect_numerals", preconds = [], 
   3.918 +      rew_ord = ("dummy_ord", dummy_ord),
   3.919 +      erls = Atools_erls(*erls3.4.03*),srls = Erls,
   3.920 +      calc = [("PLUS"  , ("op +", eval_binop "#add_")), 
   3.921 +	      ("TIMES" , ("op *", eval_binop "#mult_")),
   3.922 +	      ("POWER", ("Atools.pow", eval_binop "#power_"))
   3.923 +	      ],
   3.924 +      (*asm_thm = [],*)
   3.925 +      rules = [Thm ("real_num_collect",num_str real_num_collect), 
   3.926 +	       (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
   3.927 +	       Thm ("real_num_collect_assoc",num_str real_num_collect_assoc),
   3.928 +	       (*"[| l is_const; m is_const |] ==>  
   3.929 +				l * n + (m * n + k) =  (l + m) * n + k"*)
   3.930 +	       Thm ("real_one_collect",num_str real_one_collect),	
   3.931 +	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
   3.932 +	       Thm ("real_one_collect_assoc",num_str real_one_collect_assoc), 
   3.933 +	       (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
   3.934 +	       Calc ("op +", eval_binop "#add_"), 
   3.935 +	       Calc ("op *", eval_binop "#mult_"),
   3.936 +	       Calc ("Atools.pow", eval_binop "#power_")
   3.937 +	       ], scr = EmptyScr}:rls;
   3.938 +val reduce_012 = 
   3.939 +  Rls{id = "reduce_012", preconds = [], 
   3.940 +      rew_ord = ("dummy_ord", dummy_ord),
   3.941 +      erls = e_rls,srls = Erls,
   3.942 +      calc = [],
   3.943 +      (*asm_thm = [],*)
   3.944 +      rules = [Thm ("real_mult_1",num_str real_mult_1),                 
   3.945 +	       (*"1 * z = z"*)
   3.946 +	       (*Thm ("real_mult_minus1",num_str real_mult_minus1),14.3.03*)
   3.947 +	       (*"-1 * z = - z"*)
   3.948 +	       Thm ("sym_real_mult_minus_eq1", 
   3.949 +		    num_str (real_mult_minus_eq1 RS sym)),
   3.950 +	       (*- (?x * ?y) = "- ?x * ?y"*)
   3.951 +	       (*Thm ("real_minus_mult_cancel",num_str real_minus_mult_cancel),
   3.952 +	       (*"- ?x * - ?y = ?x * ?y"*)---*)
   3.953 +	       Thm ("real_mult_0",num_str real_mult_0),        
   3.954 +	       (*"0 * z = 0"*)
   3.955 +	       Thm ("real_add_zero_left",num_str real_add_zero_left),
   3.956 +	       (*"0 + z = z"*)
   3.957 +	       Thm ("real_add_minus",num_str real_add_minus),
   3.958 +	       (*"?z + - ?z = 0"*)
   3.959 +	       Thm ("sym_real_mult_2",num_str (real_mult_2 RS sym)),	
   3.960 +	       (*"z1 + z1 = 2 * z1"*)
   3.961 +	       Thm ("real_mult_2_assoc",num_str real_mult_2_assoc)
   3.962 +	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
   3.963 +	       ], scr = EmptyScr}:rls;
   3.964 +(*ein Hilfs-'ruleset' (benutzt das leere 'ruleset')*)
   3.965 +val discard_parentheses = 
   3.966 +    append_rls "discard_parentheses" e_rls 
   3.967 +	       [Thm ("sym_real_mult_assoc", num_str (real_mult_assoc RS sym)),
   3.968 +		Thm ("sym_real_add_assoc",num_str (real_add_assoc RS sym))];
   3.969 +
   3.970 +val scr_make_polynomial = 
   3.971 +"Script Expand_binoms t_ =                                  " ^
   3.972 +"(Repeat                                                    " ^
   3.973 +"((Try (Repeat (Rewrite real_diff_minus         False))) @@ " ^ 
   3.974 +
   3.975 +" (Try (Repeat (Rewrite real_add_mult_distrib   False))) @@ " ^	 
   3.976 +" (Try (Repeat (Rewrite real_add_mult_distrib2  False))) @@ " ^	
   3.977 +" (Try (Repeat (Rewrite real_diff_mult_distrib  False))) @@ " ^	
   3.978 +" (Try (Repeat (Rewrite real_diff_mult_distrib2 False))) @@ " ^	
   3.979 +
   3.980 +" (Try (Repeat (Rewrite real_mult_1             False))) @@ " ^		   
   3.981 +" (Try (Repeat (Rewrite real_mult_0             False))) @@ " ^		   
   3.982 +" (Try (Repeat (Rewrite real_add_zero_left      False))) @@ " ^	 
   3.983 +
   3.984 +" (Try (Repeat (Rewrite real_mult_commute       False))) @@ " ^		
   3.985 +" (Try (Repeat (Rewrite real_mult_left_commute  False))) @@ " ^	
   3.986 +" (Try (Repeat (Rewrite real_mult_assoc         False))) @@ " ^		
   3.987 +" (Try (Repeat (Rewrite real_add_commute        False))) @@ " ^		
   3.988 +" (Try (Repeat (Rewrite real_add_left_commute   False))) @@ " ^	 
   3.989 +" (Try (Repeat (Rewrite real_add_assoc          False))) @@ " ^	 
   3.990 +
   3.991 +" (Try (Repeat (Rewrite sym_realpow_twoI        False))) @@ " ^	 
   3.992 +" (Try (Repeat (Rewrite realpow_plus_1          False))) @@ " ^	 
   3.993 +" (Try (Repeat (Rewrite sym_real_mult_2         False))) @@ " ^		
   3.994 +" (Try (Repeat (Rewrite real_mult_2_assoc       False))) @@ " ^		
   3.995 +
   3.996 +" (Try (Repeat (Rewrite real_num_collect        False))) @@ " ^		
   3.997 +" (Try (Repeat (Rewrite real_num_collect_assoc  False))) @@ " ^	
   3.998 +
   3.999 +" (Try (Repeat (Rewrite real_one_collect        False))) @@ " ^		
  3.1000 +" (Try (Repeat (Rewrite real_one_collect_assoc  False))) @@ " ^   
  3.1001 +
  3.1002 +" (Try (Repeat (Calculate plus  ))) @@                      " ^
  3.1003 +" (Try (Repeat (Calculate times ))) @@                      " ^
  3.1004 +" (Try (Repeat (Calculate power_))))                        " ^  
  3.1005 +" t_)";
  3.1006 +
  3.1007 +(*version used by MG.02/03, overwritten by version AG in 04 below
  3.1008 +val make_polynomial = prep_rls(
  3.1009 +  Seq{id = "make_polynomial", preconds = []:term list, 
  3.1010 +      rew_ord = ("dummy_ord", dummy_ord),
  3.1011 +      erls = Atools_erls, srls = Erls,
  3.1012 +      calc = [],(*asm_thm = [],*)
  3.1013 +      rules = [Rls_ expand_poly,
  3.1014 +	       Rls_ order_add_mult,
  3.1015 +	       Rls_ simplify_power,   (*realpow_eq_oneI, eg. x^1 --> x *)
  3.1016 +	       Rls_ collect_numerals, (*eg. x^(2+ -1) --> x^1          *)
  3.1017 +	       Rls_ reduce_012,
  3.1018 +	       Thm ("realpow_oneI",num_str realpow_oneI),(*in --^*) 
  3.1019 +	       Rls_ discard_parentheses
  3.1020 +	       ],
  3.1021 +      scr = EmptyScr
  3.1022 +      }:rls);   *)
  3.1023 +
  3.1024 +val scr_expand_binoms =
  3.1025 +"Script Expand_binoms t_ =" ^
  3.1026 +"(Repeat                       " ^
  3.1027 +"((Try (Repeat (Rewrite real_plus_binom_pow2    False))) @@ " ^
  3.1028 +" (Try (Repeat (Rewrite real_plus_binom_times   False))) @@ " ^
  3.1029 +" (Try (Repeat (Rewrite real_minus_binom_pow2   False))) @@ " ^
  3.1030 +" (Try (Repeat (Rewrite real_minus_binom_times  False))) @@ " ^
  3.1031 +" (Try (Repeat (Rewrite real_plus_minus_binom1  False))) @@ " ^
  3.1032 +" (Try (Repeat (Rewrite real_plus_minus_binom2  False))) @@ " ^
  3.1033 +
  3.1034 +" (Try (Repeat (Rewrite real_mult_1             False))) @@ " ^
  3.1035 +" (Try (Repeat (Rewrite real_mult_0             False))) @@ " ^
  3.1036 +" (Try (Repeat (Rewrite real_add_zero_left      False))) @@ " ^
  3.1037 +
  3.1038 +" (Try (Repeat (Calculate plus  ))) @@ " ^
  3.1039 +" (Try (Repeat (Calculate times ))) @@ " ^
  3.1040 +" (Try (Repeat (Calculate power_))) @@ " ^
  3.1041 +
  3.1042 +" (Try (Repeat (Rewrite sym_realpow_twoI        False))) @@ " ^
  3.1043 +" (Try (Repeat (Rewrite realpow_plus_1          False))) @@ " ^
  3.1044 +" (Try (Repeat (Rewrite sym_real_mult_2         False))) @@ " ^
  3.1045 +" (Try (Repeat (Rewrite real_mult_2_assoc       False))) @@ " ^
  3.1046 +
  3.1047 +" (Try (Repeat (Rewrite real_num_collect        False))) @@ " ^
  3.1048 +" (Try (Repeat (Rewrite real_num_collect_assoc  False))) @@ " ^
  3.1049 +
  3.1050 +" (Try (Repeat (Rewrite real_one_collect        False))) @@ " ^
  3.1051 +" (Try (Repeat (Rewrite real_one_collect_assoc  False))) @@ " ^ 
  3.1052 +
  3.1053 +" (Try (Repeat (Calculate plus  ))) @@ " ^
  3.1054 +" (Try (Repeat (Calculate times ))) @@ " ^
  3.1055 +" (Try (Repeat (Calculate power_)))) " ^  
  3.1056 +" t_)";
  3.1057 +
  3.1058 +val expand_binoms = 
  3.1059 +  Rls{id = "expand_binoms", preconds = [], rew_ord = ("termlessI",termlessI),
  3.1060 +      erls = Atools_erls, srls = Erls,
  3.1061 +      calc = [("PLUS"  , ("op +", eval_binop "#add_")), 
  3.1062 +	      ("TIMES" , ("op *", eval_binop "#mult_")),
  3.1063 +	      ("POWER", ("Atools.pow", eval_binop "#power_"))
  3.1064 +	      ],
  3.1065 +      (*asm_thm = [],*)
  3.1066 +      rules = [Thm ("real_plus_binom_pow2"  ,num_str real_plus_binom_pow2),     
  3.1067 +	       (*"(a + b) ^^^ 2 = a ^^^ 2 + 2 * a * b + b ^^^ 2"*)
  3.1068 +	       Thm ("real_plus_binom_times" ,num_str real_plus_binom_times),    
  3.1069 +	      (*"(a + b)*(a + b) = ...*)
  3.1070 +	       Thm ("real_minus_binom_pow2" ,num_str real_minus_binom_pow2),   
  3.1071 +	       (*"(a - b) ^^^ 2 = a ^^^ 2 - 2 * a * b + b ^^^ 2"*)
  3.1072 +	       Thm ("real_minus_binom_times",num_str real_minus_binom_times),   
  3.1073 +	       (*"(a - b)*(a - b) = ...*)
  3.1074 +	       Thm ("real_plus_minus_binom1",num_str real_plus_minus_binom1),   
  3.1075 +		(*"(a + b) * (a - b) = a ^^^ 2 - b ^^^ 2"*)
  3.1076 +	       Thm ("real_plus_minus_binom2",num_str real_plus_minus_binom2),   
  3.1077 +		(*"(a - b) * (a + b) = a ^^^ 2 - b ^^^ 2"*)
  3.1078 +	       (*RL 020915*)
  3.1079 +	       Thm ("real_pp_binom_times",num_str real_pp_binom_times), 
  3.1080 +		(*(a + b)*(c + d) = a*c + a*d + b*c + b*d*)
  3.1081 +               Thm ("real_pm_binom_times",num_str real_pm_binom_times), 
  3.1082 +		(*(a + b)*(c - d) = a*c - a*d + b*c - b*d*)
  3.1083 +               Thm ("real_mp_binom_times",num_str real_mp_binom_times), 
  3.1084 +		(*(a - b)*(c + d) = a*c + a*d - b*c - b*d*)
  3.1085 +               Thm ("real_mm_binom_times",num_str real_mm_binom_times), 
  3.1086 +		(*(a - b)*(c - d) = a*c - a*d - b*c + b*d*)
  3.1087 +	       Thm ("realpow_multI",num_str realpow_multI),                
  3.1088 +		(*(a*b)^^^n = a^^^n * b^^^n*)
  3.1089 +	       Thm ("real_plus_binom_pow3",num_str real_plus_binom_pow3),
  3.1090 +	        (* (a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3 *)
  3.1091 +	       Thm ("real_minus_binom_pow3",num_str real_minus_binom_pow3),
  3.1092 +	        (* (a - b)^^^3 = a^^^3 - 3*a^^^2*b + 3*a*b^^^2 - b^^^3 *)
  3.1093 +
  3.1094 +
  3.1095 +             (*  Thm ("real_add_mult_distrib" ,num_str real_add_mult_distrib),	
  3.1096 +		(*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
  3.1097 +	       Thm ("real_add_mult_distrib2",num_str real_add_mult_distrib2),	
  3.1098 +	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
  3.1099 +	       Thm ("real_diff_mult_distrib" ,num_str real_diff_mult_distrib),	
  3.1100 +	       (*"(z1.0 - z2.0) * w = z1.0 * w - z2.0 * w"*)
  3.1101 +	       Thm ("real_diff_mult_distrib2",num_str real_diff_mult_distrib2),	
  3.1102 +	       (*"w * (z1.0 - z2.0) = w * z1.0 - w * z2.0"*)
  3.1103 +	       *)
  3.1104 +	       
  3.1105 +	       Thm ("real_mult_1",num_str real_mult_1),              (*"1 * z = z"*)
  3.1106 +	       Thm ("real_mult_0",num_str real_mult_0),              (*"0 * z = 0"*)
  3.1107 +	       Thm ("real_add_zero_left",num_str real_add_zero_left),(*"0 + z = z"*)
  3.1108 +
  3.1109 +	       Calc ("op +", eval_binop "#add_"), 
  3.1110 +	       Calc ("op *", eval_binop "#mult_"),
  3.1111 +	       Calc ("Atools.pow", eval_binop "#power_"),
  3.1112 +               (*	       
  3.1113 +	        Thm ("real_mult_commute",num_str real_mult_commute),		(*AC-rewriting*)
  3.1114 +	       Thm ("real_mult_left_commute",num_str real_mult_left_commute),	(**)
  3.1115 +	       Thm ("real_mult_assoc",num_str real_mult_assoc),			(**)
  3.1116 +	       Thm ("real_add_commute",num_str real_add_commute),		(**)
  3.1117 +	       Thm ("real_add_left_commute",num_str real_add_left_commute),	(**)
  3.1118 +	       Thm ("real_add_assoc",num_str real_add_assoc),	                (**)
  3.1119 +	       *)
  3.1120 +	       
  3.1121 +	       Thm ("sym_realpow_twoI",num_str (realpow_twoI RS sym)),		
  3.1122 +	       (*"r1 * r1 = r1 ^^^ 2"*)
  3.1123 +	       Thm ("realpow_plus_1",num_str realpow_plus_1),			
  3.1124 +	       (*"r * r ^^^ n = r ^^^ (n + 1)"*)
  3.1125 +	       (*Thm ("sym_real_mult_2",num_str (real_mult_2 RS sym)),		
  3.1126 +	       (*"z1 + z1 = 2 * z1"*)*)
  3.1127 +	       Thm ("real_mult_2_assoc",num_str real_mult_2_assoc),		
  3.1128 +	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
  3.1129 +
  3.1130 +	       Thm ("real_num_collect",num_str real_num_collect), 
  3.1131 +	       (*"[| l is_const; m is_const |] ==> l * n + m * n = (l + m) * n"*)
  3.1132 +	       Thm ("real_num_collect_assoc",num_str real_num_collect_assoc),	
  3.1133 +	       (*"[| l is_const; m is_const |] ==>  l * n + (m * n + k) =  (l + m) * n + k"*)
  3.1134 +	       Thm ("real_one_collect",num_str real_one_collect),		
  3.1135 +	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
  3.1136 +	       Thm ("real_one_collect_assoc",num_str real_one_collect_assoc), 
  3.1137 +	       (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
  3.1138 +
  3.1139 +	       Calc ("op +", eval_binop "#add_"), 
  3.1140 +	       Calc ("op *", eval_binop "#mult_"),
  3.1141 +	       Calc ("Atools.pow", eval_binop "#power_")
  3.1142 +	       ],
  3.1143 +      scr = Script ((term_of o the o (parse thy)) scr_expand_binoms)
  3.1144 +      }:rls;      
  3.1145 +
  3.1146 +
  3.1147 +"******* Poly.ML end ******* ...RL";
  3.1148 +
  3.1149 +
  3.1150 +(**. MG.03: make_polynomial_ ... uses SML-fun for ordering .**)
  3.1151 +
  3.1152 +(*FIXME.0401: make SML-order local to make_polynomial(_) *)
  3.1153 +(*FIXME.0401: replace 'make_polynomial'(old) by 'make_polynomial_'(MG) *)
  3.1154 +(* Polynom --> List von Monomen *) 
  3.1155 +fun poly2list (Const ("op +",_) $ t1 $ t2) = 
  3.1156 +    (poly2list t1) @ (poly2list t2)
  3.1157 +  | poly2list t = [t];
  3.1158 +
  3.1159 +(* Monom --> Liste von Variablen *)
  3.1160 +fun monom2list (Const ("op *",_) $ t1 $ t2) = 
  3.1161 +    (monom2list t1) @ (monom2list t2)
  3.1162 +  | monom2list t = [t];
  3.1163 +
  3.1164 +(* liefert Variablenname (String) einer Variablen und Basis bei Potenz *)
  3.1165 +fun get_basStr (Const ("Atools.pow",_) $ Free (str, _) $ _) = str
  3.1166 +  | get_basStr (Free (str, _)) = str
  3.1167 +  | get_basStr t = "|||"; (* gross gewichtet; für Brüch ect. *)
  3.1168 +(*| get_basStr t = 
  3.1169 +    raise error("get_basStr: called with t= "^(term2str t));*)
  3.1170 +
  3.1171 +(* liefert Hochzahl (String) einer Variablen bzw Gewichtstring (zum Sortieren) *)
  3.1172 +fun get_potStr (Const ("Atools.pow",_) $ Free _ $ Free (str, _)) = str
  3.1173 +  | get_potStr (Const ("Atools.pow",_) $ Free _ $ _ ) = "|||" (* gross gewichtet *)
  3.1174 +  | get_potStr (Free (str, _)) = "---" (* keine Hochzahl --> kleinst gewichtet *)
  3.1175 +  | get_potStr t = "||||||"; (* gross gewichtet; für Brüch ect. *)
  3.1176 +(*| get_potStr t = 
  3.1177 +    raise error("get_potStr: called with t= "^(term2str t));*)
  3.1178 +
  3.1179 +(* Umgekehrte string_ord *)
  3.1180 +val string_ord_rev =  rev_order o string_ord;
  3.1181 +		
  3.1182 + (* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen) 
  3.1183 +    innerhalb eines Monomes:
  3.1184 +    - zuerst lexikographisch nach Variablenname 
  3.1185 +    - wenn gleich: nach steigender Potenz *)
  3.1186 +fun var_ord (a,b: term) = prod_ord string_ord string_ord 
  3.1187 +    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
  3.1188 +
  3.1189 +(* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen); 
  3.1190 +   verwendet zum Sortieren von Monomen mittels Gesamtgradordnung:
  3.1191 +   - zuerst lexikographisch nach Variablenname 
  3.1192 +   - wenn gleich: nach sinkender Potenz*)
  3.1193 +fun var_ord_revPow (a,b: term) = prod_ord string_ord string_ord_rev 
  3.1194 +    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
  3.1195 +
  3.1196 +
  3.1197 +(* Ordnet ein Liste von Variablen (und Potenzen) lexikographisch *)
  3.1198 +val sort_varList = sort var_ord;
  3.1199 +
  3.1200 +(* Entfernet aeussersten Operator (Wurzel) aus einem Term und schreibt 
  3.1201 +   Argumente in eine Liste *)
  3.1202 +fun args u : term list =
  3.1203 +    let fun stripc (f$t, ts) = stripc (f, t::ts)
  3.1204 +	  | stripc (t as Free _, ts) = (t::ts)
  3.1205 +	  | stripc (_, ts) = ts
  3.1206 +    in stripc (u, []) end;
  3.1207 +                                    
  3.1208 +(* liefert True, falls der Term (Liste von Termen) nur Zahlen 
  3.1209 +   (keine Variablen) enthaelt *)
  3.1210 +fun filter_num [] = true
  3.1211 +  | filter_num [Free x] = if (is_num (Free x)) then true
  3.1212 +				else false
  3.1213 +  | filter_num ((Free _)::_) = false
  3.1214 +  | filter_num ts =
  3.1215 +    (filter_num o (filter_out is_num) o flat o (map args)) ts;
  3.1216 +
  3.1217 +(* liefert True, falls der Term nur Zahlen (keine Variablen) enthaelt 
  3.1218 +   dh. er ist ein numerischer Wert und entspricht einem Koeffizienten *)
  3.1219 +fun is_nums t = filter_num [t];
  3.1220 +
  3.1221 +(* Berechnet den Gesamtgrad eines Monoms *)
  3.1222 +local 
  3.1223 +    fun counter (n, []) = n
  3.1224 +      | counter (n, x :: xs) = 
  3.1225 +	if (is_nums x) then
  3.1226 +	    counter (n, xs) 
  3.1227 +	else 
  3.1228 +	    (case x of 
  3.1229 +		 (Const ("Atools.pow", _) $ Free (str_b, _) $ Free (str_h, T)) => 
  3.1230 +		     if (is_nums (Free (str_h, T))) then
  3.1231 +			 counter (n + (the (int_of_str str_h)), xs)
  3.1232 +		     else counter (n + 1000, xs) (*FIXME.MG?!*)
  3.1233 +	       | (Const ("Atools.pow", _) $ Free (str_b, _) $ _ ) => 
  3.1234 +		     counter (n + 1000, xs) (*FIXME.MG?!*)
  3.1235 +	       | (Free (str, _)) => counter (n + 1, xs)
  3.1236 +	     (*| _ => raise error("monom_degree: called with factor: "^(term2str x)))*)
  3.1237 +	       | _ => counter (n + 10000, xs)) (*FIXME.MG?! ... Brüche ect.*)
  3.1238 +in  
  3.1239 +    fun monom_degree l = counter (0, l) 
  3.1240 +end;
  3.1241 +
  3.1242 +(* wie Ordnung dict_ord (lexicographische Ordnung zweier Listen, mit Vergleich 
  3.1243 +   der Listen-Elemente mit elem_ord) - Elemente die Bedingung cond erfuellen, 
  3.1244 +   werden jedoch dabei ignoriert (uebersprungen)  *)
  3.1245 +fun dict_cond_ord _ _ ([], []) = EQUAL
  3.1246 +  | dict_cond_ord _ _ ([], _ :: _) = LESS
  3.1247 +  | dict_cond_ord _ _ (_ :: _, []) = GREATER
  3.1248 +  | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
  3.1249 +    (case (cond x, cond y) of 
  3.1250 +	 (false, false) => (case elem_ord (x, y) of 
  3.1251 +				EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
  3.1252 +			      | ord => ord)
  3.1253 +       | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
  3.1254 +       | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
  3.1255 +       | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
  3.1256 +
  3.1257 +(* Gesamtgradordnung zum Vergleich von Monomen (Liste von Variablen/Potenzen):
  3.1258 +   zuerst nach Gesamtgrad, bei gleichem Gesamtgrad lexikographisch ordnen - 
  3.1259 +   dabei werden Koeffizienten ignoriert (2*3*a^^^2*4*b gilt wie a^^^2*b) *)
  3.1260 +fun degree_ord (xs, ys) =
  3.1261 +	    prod_ord int_ord (dict_cond_ord var_ord_revPow is_nums) 
  3.1262 +	    ((monom_degree xs, xs), (monom_degree ys, ys));
  3.1263 +
  3.1264 +fun hd_str str = substring (str, 0, 1);
  3.1265 +fun tl_str str = substring (str, 1, (size str) - 1);
  3.1266 +
  3.1267 +(* liefert nummerischen Koeffizienten eines Monoms oder NONE *)
  3.1268 +fun get_koeff_of_mon [] =  raise error("get_koeff_of_mon: called with l = []")
  3.1269 +  | get_koeff_of_mon (l as x::xs) = if is_nums x then SOME x
  3.1270 +				    else NONE;
  3.1271 +
  3.1272 +(* wandelt Koeffizient in (zum sortieren geeigneten) String um *)
  3.1273 +fun koeff2ordStr (SOME x) = (case x of 
  3.1274 +				 (Free (str, T)) => 
  3.1275 +				     if (hd_str str) = "-" then (tl_str str)^"0" (* 3 < -3 *)
  3.1276 +				     else str
  3.1277 +			       | _ => "aaa") (* "num.Ausdruck" --> gross *)
  3.1278 +  | koeff2ordStr NONE = "---"; (* "kein Koeff" --> kleinste *)
  3.1279 +
  3.1280 +(* Order zum Vergleich von Koeffizienten (strings): 
  3.1281 +   "kein Koeff" < "0" < "1" < "-1" < "2" < "-2" < ... < "num.Ausdruck" *)
  3.1282 +fun compare_koeff_ord (xs, ys) = 
  3.1283 +    string_ord ((koeff2ordStr o get_koeff_of_mon) xs,
  3.1284 +		(koeff2ordStr o get_koeff_of_mon) ys);
  3.1285 +
  3.1286 +(* Gesamtgradordnung degree_ord + Ordnen nach Koeffizienten falls EQUAL *)
  3.1287 +fun koeff_degree_ord (xs, ys) =
  3.1288 +	    prod_ord degree_ord compare_koeff_ord ((xs, xs), (ys, ys));
  3.1289 +
  3.1290 +(* Ordnet ein Liste von Monomen (Monom = Liste von Variablen) mittels 
  3.1291 +   Gesamtgradordnung *)
  3.1292 +val sort_monList = sort koeff_degree_ord;
  3.1293 +
  3.1294 +(* Alternativ zu degree_ord koennte auch die viel einfachere und 
  3.1295 +   kuerzere Ordnung simple_ord verwendet werden - ist aber nicht 
  3.1296 +   fuer unsere Zwecke geeignet!
  3.1297 +
  3.1298 +fun simple_ord (al,bl: term list) = dict_ord string_ord 
  3.1299 +	 (map get_basStr al, map get_basStr bl); 
  3.1300 +
  3.1301 +val sort_monList = sort simple_ord; *)
  3.1302 +
  3.1303 +(* aus 2 Variablen wird eine Summe bzw ein Produkt erzeugt 
  3.1304 +   (mit gewuenschtem Typen T) *)
  3.1305 +fun plus T = Const ("op +", [T,T] ---> T);
  3.1306 +fun mult T = Const ("op *", [T,T] ---> T);
  3.1307 +fun binop op_ t1 t2 = op_ $ t1 $ t2;
  3.1308 +fun create_prod T (a,b) = binop (mult T) a b;
  3.1309 +fun create_sum T (a,b) = binop (plus T) a b;
  3.1310 +
  3.1311 +(* löscht letztes Element einer Liste *)
  3.1312 +fun drop_last l = take ((length l)-1,l);
  3.1313 +
  3.1314 +(* Liste von Variablen --> Monom *)
  3.1315 +fun create_monom T vl = foldr (create_prod T) (drop_last vl, last_elem vl);
  3.1316 +(* Bemerkung: 
  3.1317 +   foldr bewirkt rechtslastige Klammerung des Monoms - ist notwendig, damit zwei 
  3.1318 +   gleiche Monome zusammengefasst werden können (collect_numerals)! 
  3.1319 +   zB: 2*(x*(y*z)) + 3*(x*(y*z)) --> (2+3)*(x*(y*z))*)
  3.1320 +
  3.1321 +(* Liste von Monomen --> Polynom *)	
  3.1322 +fun create_polynom T ml = foldl (create_sum T) (hd ml, tl ml);
  3.1323 +(* Bemerkung: 
  3.1324 +   foldl bewirkt linkslastige Klammerung des Polynoms (der Summanten) - 
  3.1325 +   bessere Darstellung, da keine Klammern sichtbar! 
  3.1326 +   (und discard_parentheses in make_polynomial hat weniger zu tun) *)
  3.1327 +
  3.1328 +(* sorts the variables (faktors) of an expanded polynomial lexicographical *)
  3.1329 +fun sort_variables t = 
  3.1330 +    let
  3.1331 +	val ll =  map monom2list (poly2list t);
  3.1332 +	val lls = map sort_varList ll; 
  3.1333 +	val T = type_of t;
  3.1334 +	val ls = map (create_monom T) lls;
  3.1335 +    in create_polynom T ls end;
  3.1336 +
  3.1337 +(* sorts the monoms of an expanded and variable-sorted polynomial 
  3.1338 +   by total_degree *)
  3.1339 +fun sort_monoms t = 
  3.1340 +    let
  3.1341 +	val ll =  map monom2list (poly2list t);
  3.1342 +	val lls = sort_monList ll;
  3.1343 +	val T = type_of t;
  3.1344 +	val ls = map (create_monom T) lls;
  3.1345 +    in create_polynom T ls end;
  3.1346 +
  3.1347 +(* auch Klammerung muss übereinstimmen; 
  3.1348 +   sort_variables klammert Produkte rechtslastig*)
  3.1349 +fun is_multUnordered t = ((is_polyexp t) andalso not (t = sort_variables t));
  3.1350 +
  3.1351 +fun eval_is_multUnordered (thmid:string) _ 
  3.1352 +		       (t as (Const("Poly.is'_multUnordered", _) $ arg)) thy = 
  3.1353 +    if is_multUnordered arg
  3.1354 +    then SOME (mk_thmid thmid "" 
  3.1355 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
  3.1356 +	       Trueprop $ (mk_equality (t, HOLogic.true_const)))
  3.1357 +    else SOME (mk_thmid thmid "" 
  3.1358 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
  3.1359 +	       Trueprop $ (mk_equality (t, HOLogic.false_const)))
  3.1360 +  | eval_is_multUnordered _ _ _ _ = NONE; 
  3.1361 +
  3.1362 +
  3.1363 +fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
  3.1364 +    []:(rule * (term * term list)) list;
  3.1365 +fun init_state (_:term) = e_rrlsstate;
  3.1366 +fun locate_rule (_:rule list list) (_:term) (_:rule) =
  3.1367 +    ([]:(rule * (term * term list)) list);
  3.1368 +fun next_rule (_:rule list list) (_:term) = (NONE:rule option);
  3.1369 +fun normal_form t = SOME (sort_variables t,[]:term list);
  3.1370 +
  3.1371 +val order_mult_ =
  3.1372 +    Rrls {id = "order_mult_", 
  3.1373 +	  prepat = 
  3.1374 +	  [([(term_of o the o (parse thy)) "p is_multUnordered"], 
  3.1375 +	    (term_of o the o (parse thy)) "?p" )],
  3.1376 +	  rew_ord = ("dummy_ord", dummy_ord),
  3.1377 +	  erls = append_rls "e_rls-is_multUnordered" e_rls(*MG: poly_erls*)
  3.1378 +			    [Calc ("Poly.is'_multUnordered", eval_is_multUnordered "")
  3.1379 +			     ],
  3.1380 +	  calc = [("PLUS"    ,("op +"        ,eval_binop "#add_")),
  3.1381 +		  ("TIMES"   ,("op *"        ,eval_binop "#mult_")),
  3.1382 +		  ("DIVIDE" ,("HOL.divide"  ,eval_cancel "#divide_")),
  3.1383 +		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
  3.1384 +	  (*asm_thm=[],*)
  3.1385 +	  scr=Rfuns {init_state  = init_state,
  3.1386 +		     normal_form = normal_form,
  3.1387 +		     locate_rule = locate_rule,
  3.1388 +		     next_rule   = next_rule,
  3.1389 +		     attach_form = attach_form}};
  3.1390 +
  3.1391 +val order_mult_rls_ = 
  3.1392 +  Rls{id = "order_mult_rls_", preconds = [], 
  3.1393 +      rew_ord = ("dummy_ord", dummy_ord),
  3.1394 +      erls = e_rls,srls = Erls,
  3.1395 +      calc = [],
  3.1396 +      (*asm_thm = [],*)
  3.1397 +      rules = [Rls_ order_mult_
  3.1398 +	       ], scr = EmptyScr}:rls;
  3.1399 +
  3.1400 +fun is_addUnordered t = ((is_polyexp t) andalso not (t = sort_monoms t));
  3.1401 +
  3.1402 +(*WN.18.6.03 *)
  3.1403 +(*("is_addUnordered", ("Poly.is'_addUnordered", eval_is_addUnordered ""))*)
  3.1404 +fun eval_is_addUnordered (thmid:string) _ 
  3.1405 +		       (t as (Const("Poly.is'_addUnordered", _) $ arg)) thy = 
  3.1406 +    if is_addUnordered arg
  3.1407 +    then SOME (mk_thmid thmid "" 
  3.1408 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
  3.1409 +	       Trueprop $ (mk_equality (t, HOLogic.true_const)))
  3.1410 +    else SOME (mk_thmid thmid "" 
  3.1411 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
  3.1412 +	       Trueprop $ (mk_equality (t, HOLogic.false_const)))
  3.1413 +  | eval_is_addUnordered _ _ _ _ = NONE; 
  3.1414 +
  3.1415 +fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
  3.1416 +    []:(rule * (term * term list)) list;
  3.1417 +fun init_state (_:term) = e_rrlsstate;
  3.1418 +fun locate_rule (_:rule list list) (_:term) (_:rule) =
  3.1419 +    ([]:(rule * (term * term list)) list);
  3.1420 +fun next_rule (_:rule list list) (_:term) = (NONE:rule option);
  3.1421 +fun normal_form t = SOME (sort_monoms t,[]:term list);
  3.1422 +
  3.1423 +val order_add_ =
  3.1424 +    Rrls {id = "order_add_", 
  3.1425 +	  prepat = (*WN.18.6.03 Preconditions und Pattern,
  3.1426 +		    die beide passen muessen, damit das Rrls angewandt wird*)
  3.1427 +	  [([(term_of o the o (parse thy)) "p is_addUnordered"], 
  3.1428 +	    (term_of o the o (parse thy)) "?p" 
  3.1429 +	    (*WN.18.6.03 also KEIN pattern, dieses erzeugt nur das Environment 
  3.1430 +	      fuer die Evaluation der Precondition "p is_addUnordered"*))],
  3.1431 +	  rew_ord = ("dummy_ord", dummy_ord),
  3.1432 +	  erls = append_rls "e_rls-is_addUnordered" e_rls(*MG: poly_erls*)
  3.1433 +			    [Calc ("Poly.is'_addUnordered", eval_is_addUnordered "")
  3.1434 +			     (*WN.18.6.03 definiert in Poly.thy,
  3.1435 +                               evaluiert prepat*)],
  3.1436 +	  calc = [("PLUS"    ,("op +"        ,eval_binop "#add_")),
  3.1437 +		  ("TIMES"   ,("op *"        ,eval_binop "#mult_")),
  3.1438 +		  ("DIVIDE" ,("HOL.divide"  ,eval_cancel "#divide_")),
  3.1439 +		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
  3.1440 +	  (*asm_thm=[],*)
  3.1441 +	  scr=Rfuns {init_state  = init_state,
  3.1442 +		     normal_form = normal_form,
  3.1443 +		     locate_rule = locate_rule,
  3.1444 +		     next_rule   = next_rule,
  3.1445 +		     attach_form = attach_form}};
  3.1446 +
  3.1447 +val order_add_rls_ = 
  3.1448 +  Rls{id = "order_add_rls_", preconds = [], 
  3.1449 +      rew_ord = ("dummy_ord", dummy_ord),
  3.1450 +      erls = e_rls,srls = Erls,
  3.1451 +      calc = [],
  3.1452 +      (*asm_thm = [],*)
  3.1453 +      rules = [Rls_ order_add_
  3.1454 +	       ], scr = EmptyScr}:rls;
  3.1455 +
  3.1456 +(*. see MG-DA.p.52ff .*)
  3.1457 +val make_polynomial(*MG.03, overwrites version from above, 
  3.1458 +    previously 'make_polynomial_'*) =
  3.1459 +  Seq {id = "make_polynomial", preconds = []:term list, 
  3.1460 +      rew_ord = ("dummy_ord", dummy_ord),
  3.1461 +      erls = Atools_erls, srls = Erls,calc = [],
  3.1462 +      rules = [Rls_ discard_minus_,
  3.1463 +	       Rls_ expand_poly_,
  3.1464 +	       Calc ("op *", eval_binop "#mult_"),
  3.1465 +	       Rls_ order_mult_rls_,
  3.1466 +	       Rls_ simplify_power_, 
  3.1467 +	       Rls_ calc_add_mult_pow_, 
  3.1468 +	       Rls_ reduce_012_mult_,
  3.1469 +	       Rls_ order_add_rls_,
  3.1470 +	       Rls_ collect_numerals_, 
  3.1471 +	       Rls_ reduce_012_,
  3.1472 +	       Rls_ discard_parentheses_
  3.1473 +	       ],
  3.1474 +      scr = EmptyScr
  3.1475 +      }:rls;
  3.1476 +val norm_Poly(*=make_polynomial*) = 
  3.1477 +  Seq {id = "norm_Poly", preconds = []:term list, 
  3.1478 +      rew_ord = ("dummy_ord", dummy_ord),
  3.1479 +      erls = Atools_erls, srls = Erls, calc = [],
  3.1480 +      rules = [Rls_ discard_minus_,
  3.1481 +	       Rls_ expand_poly_,
  3.1482 +	       Calc ("op *", eval_binop "#mult_"),
  3.1483 +	       Rls_ order_mult_rls_,
  3.1484 +	       Rls_ simplify_power_, 
  3.1485 +	       Rls_ calc_add_mult_pow_, 
  3.1486 +	       Rls_ reduce_012_mult_,
  3.1487 +	       Rls_ order_add_rls_,
  3.1488 +	       Rls_ collect_numerals_, 
  3.1489 +	       Rls_ reduce_012_,
  3.1490 +	       Rls_ discard_parentheses_
  3.1491 +	       ],
  3.1492 +      scr = EmptyScr
  3.1493 +      }:rls;
  3.1494 +
  3.1495 +(* MG:03 Like make_polynomial_ but without Rls_ discard_parentheses_ 
  3.1496 +   and expand_poly_rat_ instead of expand_poly_, see MG-DA.p.56ff*)
  3.1497 +(* MG necessary  for termination of norm_Rational(*_mg*) in Rational.ML*)
  3.1498 +val make_rat_poly_with_parentheses =
  3.1499 +  Seq{id = "make_rat_poly_with_parentheses", preconds = []:term list, 
  3.1500 +      rew_ord = ("dummy_ord", dummy_ord),
  3.1501 +      erls = Atools_erls, srls = Erls, calc = [],
  3.1502 +      rules = [Rls_ discard_minus_,
  3.1503 +	       Rls_ expand_poly_rat_,(*ignors rationals*)
  3.1504 +	       Calc ("op *", eval_binop "#mult_"),
  3.1505 +	       Rls_ order_mult_rls_,
  3.1506 +	       Rls_ simplify_power_, 
  3.1507 +	       Rls_ calc_add_mult_pow_, 
  3.1508 +	       Rls_ reduce_012_mult_,
  3.1509 +	       Rls_ order_add_rls_,
  3.1510 +	       Rls_ collect_numerals_, 
  3.1511 +	       Rls_ reduce_012_
  3.1512 +	       (*Rls_ discard_parentheses_ *)
  3.1513 +	       ],
  3.1514 +      scr = EmptyScr
  3.1515 +      }:rls;
  3.1516 +
  3.1517 +(*.a minimal ruleset for reverse rewriting of factions [2];
  3.1518 +   compare expand_binoms.*)
  3.1519 +val rev_rew_p = 
  3.1520 +Seq{id = "reverse_rewriting", preconds = [], rew_ord = ("termlessI",termlessI),
  3.1521 +    erls = Atools_erls, srls = Erls,
  3.1522 +    calc = [(*("PLUS"  , ("op +", eval_binop "#add_")), 
  3.1523 +	    ("TIMES" , ("op *", eval_binop "#mult_")),
  3.1524 +	    ("POWER", ("Atools.pow", eval_binop "#power_"))*)
  3.1525 +	    ],
  3.1526 +    rules = [Thm ("real_plus_binom_times" ,num_str real_plus_binom_times),
  3.1527 +	     (*"(a + b)*(a + b) = a ^ 2 + 2 * a * b + b ^ 2*)
  3.1528 +	     Thm ("real_plus_binom_times1" ,num_str real_plus_binom_times1),
  3.1529 +	     (*"(a +  1*b)*(a + -1*b) = a^^^2 + -1*b^^^2"*)
  3.1530 +	     Thm ("real_plus_binom_times2" ,num_str real_plus_binom_times2),
  3.1531 +	     (*"(a + -1*b)*(a +  1*b) = a^^^2 + -1*b^^^2"*)
  3.1532 +
  3.1533 +	     Thm ("real_mult_1",num_str real_mult_1),(*"1 * z = z"*)
  3.1534 +
  3.1535 +             Thm ("real_add_mult_distrib" ,num_str real_add_mult_distrib),
  3.1536 +	     (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
  3.1537 +	     Thm ("real_add_mult_distrib2",num_str real_add_mult_distrib2),
  3.1538 +	     (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
  3.1539 +	       
  3.1540 +	     Thm ("real_mult_assoc", num_str real_mult_assoc),
  3.1541 +	     (*"?z1.1 * ?z2.1 * ?z3. =1 ?z1.1 * (?z2.1 * ?z3.1)"*)
  3.1542 +	     Rls_ order_mult_rls_,
  3.1543 +	     (*Rls_ order_add_rls_,*)
  3.1544 +
  3.1545 +	     Calc ("op +", eval_binop "#add_"), 
  3.1546 +	     Calc ("op *", eval_binop "#mult_"),
  3.1547 +	     Calc ("Atools.pow", eval_binop "#power_"),
  3.1548 +	     
  3.1549 +	     Thm ("sym_realpow_twoI",num_str (realpow_twoI RS sym)),
  3.1550 +	     (*"r1 * r1 = r1 ^^^ 2"*)
  3.1551 +	     Thm ("sym_real_mult_2",num_str (real_mult_2 RS sym)),
  3.1552 +	     (*"z1 + z1 = 2 * z1"*)
  3.1553 +	     Thm ("real_mult_2_assoc",num_str real_mult_2_assoc),
  3.1554 +	     (*"z1 + (z1 + k) = 2 * z1 + k"*)
  3.1555 +
  3.1556 +	     Thm ("real_num_collect",num_str real_num_collect), 
  3.1557 +	     (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
  3.1558 +	     Thm ("real_num_collect_assoc",num_str real_num_collect_assoc),
  3.1559 +	     (*"[| l is_const; m is_const |] ==>  
  3.1560 +                                     l * n + (m * n + k) =  (l + m) * n + k"*)
  3.1561 +	     Thm ("real_one_collect",num_str real_one_collect),
  3.1562 +	     (*"m is_const ==> n + m * n = (1 + m) * n"*)
  3.1563 +	     Thm ("real_one_collect_assoc",num_str real_one_collect_assoc), 
  3.1564 +	     (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
  3.1565 +
  3.1566 +	     Thm ("realpow_multI", num_str realpow_multI),
  3.1567 +	     (*"(r * s) ^^^ n = r ^^^ n * s ^^^ n"*)
  3.1568 +
  3.1569 +	     Calc ("op +", eval_binop "#add_"), 
  3.1570 +	     Calc ("op *", eval_binop "#mult_"),
  3.1571 +	     Calc ("Atools.pow", eval_binop "#power_"),
  3.1572 +
  3.1573 +	     Thm ("real_mult_1",num_str real_mult_1),(*"1 * z = z"*)
  3.1574 +	     Thm ("real_mult_0",num_str real_mult_0),(*"0 * z = 0"*)
  3.1575 +	     Thm ("real_add_zero_left",num_str real_add_zero_left)(*0 + z = z*)
  3.1576 +
  3.1577 +	     (*Rls_ order_add_rls_*)
  3.1578 +	     ],
  3.1579 +
  3.1580 +    scr = EmptyScr}:rls;      
  3.1581 +
  3.1582 +ruleset' := 
  3.1583 +overwritelthy thy (!ruleset',
  3.1584 +		   [("norm_Poly", prep_rls norm_Poly),
  3.1585 +		    ("Poly_erls",Poly_erls)(*FIXXXME:del with rls.rls'*),
  3.1586 +		    ("expand", prep_rls expand),
  3.1587 +		    ("expand_poly", prep_rls expand_poly),
  3.1588 +		    ("simplify_power", prep_rls simplify_power),
  3.1589 +		    ("order_add_mult", prep_rls order_add_mult),
  3.1590 +		    ("collect_numerals", prep_rls collect_numerals),
  3.1591 +		    ("collect_numerals_", prep_rls collect_numerals_),
  3.1592 +		    ("reduce_012", prep_rls reduce_012),
  3.1593 +		    ("discard_parentheses", prep_rls discard_parentheses),
  3.1594 +		    ("make_polynomial", prep_rls make_polynomial),
  3.1595 +		    ("expand_binoms", prep_rls expand_binoms),
  3.1596 +		    ("rev_rew_p", prep_rls rev_rew_p),
  3.1597 +		    ("discard_minus_", prep_rls discard_minus_),
  3.1598 +		    ("expand_poly_", prep_rls expand_poly_),
  3.1599 +		    ("expand_poly_rat_", prep_rls expand_poly_rat_),
  3.1600 +		    ("simplify_power_", prep_rls simplify_power_),
  3.1601 +		    ("calc_add_mult_pow_", prep_rls calc_add_mult_pow_),
  3.1602 +		    ("reduce_012_mult_", prep_rls reduce_012_mult_),
  3.1603 +		    ("reduce_012_", prep_rls reduce_012_),
  3.1604 +		    ("discard_parentheses_",prep_rls discard_parentheses_),
  3.1605 +		    ("order_mult_rls_", prep_rls order_mult_rls_),
  3.1606 +		    ("order_add_rls_", prep_rls order_add_rls_),
  3.1607 +		    ("make_rat_poly_with_parentheses", 
  3.1608 +		     prep_rls make_rat_poly_with_parentheses)
  3.1609 +		    (*("", prep_rls ),
  3.1610 +		     ("", prep_rls ),
  3.1611 +		     ("", prep_rls )
  3.1612 +		     *)
  3.1613 +		    ]);
  3.1614 +
  3.1615 +calclist':= overwritel (!calclist', 
  3.1616 +   [("is_polyrat_in", ("Poly.is'_polyrat'_in", 
  3.1617 +		       eval_is_polyrat_in "#eval_is_polyrat_in")),
  3.1618 +    ("is_expanded_in", ("Poly.is'_expanded'_in", eval_is_expanded_in "")),
  3.1619 +    ("is_poly_in", ("Poly.is'_poly'_in", eval_is_poly_in "")),
  3.1620 +    ("has_degree_in", ("Poly.has'_degree'_in", eval_has_degree_in "")),
  3.1621 +    ("is_polyexp", ("Poly.is'_polyexp", eval_is_polyexp "")),
  3.1622 +    ("is_multUnordered", ("Poly.is'_multUnordered", eval_is_multUnordered"")),
  3.1623 +    ("is_addUnordered", ("Poly.is'_addUnordered", eval_is_addUnordered ""))
  3.1624 +    ]);
  3.1625 +
  3.1626 +
  3.1627 +(** problems **)
  3.1628 +
  3.1629 +store_pbt
  3.1630 + (prep_pbt Poly.thy "pbl_simp_poly" [] e_pblID
  3.1631 + (["polynomial","simplification"],
  3.1632 +  [("#Given" ,["term t_"]),
  3.1633 +   ("#Where" ,["t_ is_polyexp"]),
  3.1634 +   ("#Find"  ,["normalform n_"])
  3.1635 +  ],
  3.1636 +  append_rls "e_rls" e_rls [(*for preds in where_*)
  3.1637 +			    Calc ("Poly.is'_polyexp", eval_is_polyexp "")], 
  3.1638 +  SOME "Simplify t_", 
  3.1639 +  [["simplification","for_polynomials"]]));
  3.1640 +
  3.1641 +
  3.1642 +(** methods **)
  3.1643 +
  3.1644 +store_met
  3.1645 +    (prep_met Poly.thy "met_simp_poly" [] e_metID
  3.1646 +	      (["simplification","for_polynomials"],
  3.1647 +	       [("#Given" ,["term t_"]),
  3.1648 +		("#Where" ,["t_ is_polyexp"]),
  3.1649 +		("#Find"  ,["normalform n_"])
  3.1650 +		],
  3.1651 +	       {rew_ord'="tless_true",
  3.1652 +		rls' = e_rls,
  3.1653 +		calc = [], 
  3.1654 +		srls = e_rls, 
  3.1655 +		prls = append_rls "simplification_for_polynomials_prls" e_rls 
  3.1656 +				  [(*for preds in where_*)
  3.1657 +				   Calc ("Poly.is'_polyexp",eval_is_polyexp"")],
  3.1658 +		crls = e_rls, nrls = norm_Poly},
  3.1659 +	       "Script SimplifyScript (t_::real) =                " ^
  3.1660 +	       "  ((Rewrite_Set norm_Poly False) t_)"
  3.1661 +	       ));
  3.1662 +*}
  3.1663 +
  3.1664  end
     4.1 --- a/src/Tools/isac/Knowledge/PolyMinus.thy	Thu Aug 26 10:03:53 2010 +0200
     4.2 +++ b/src/Tools/isac/Knowledge/PolyMinus.thy	Thu Aug 26 18:15:30 2010 +0200
     4.3 @@ -1,12 +1,9 @@
     4.4  (* attempts to perserve binary minus as wanted by Austrian teachers
     4.5     WN071207
     4.6     (c) due to copyright terms
     4.7 -remove_thy"PolyMinus";
     4.8 -use_thy_only"Knowledge/PolyMinus";
     4.9 -use_thy"Knowledge/Isac";
    4.10  *)
    4.11  
    4.12 -PolyMinus = (*Poly// due to "is_ratpolyexp" in...*) Rational + 
    4.13 +theory PolyMinus imports (*Poly// due to "is_ratpolyexp" in...*) Rational begin
    4.14  
    4.15  consts
    4.16  
    4.17 @@ -19,13 +16,13 @@
    4.18  	(*"Probe (3*a+2*b+a = 4*a+2*b) [a=1,b=2]"*)
    4.19  
    4.20    (*descriptions for the pbl and met*)
    4.21 -  Pruefe      :: bool => una
    4.22 -  mitWert     :: bool list => tobooll
    4.23 -  Geprueft    :: bool => una
    4.24 +  Pruefe      :: "bool => una"
    4.25 +  mitWert     :: "bool list => tobooll"
    4.26 +  Geprueft    :: "bool => una"
    4.27  
    4.28    (*Script-name*)
    4.29 -  ProbeScript :: "[bool, bool list,       bool] \
    4.30 -				      \=> bool"
    4.31 +  ProbeScript :: "[bool, bool list,       bool] 
    4.32 +				      => bool"
    4.33                    ("((Script ProbeScript (_ _ =))// (_))" 9)
    4.34  
    4.35  rules
    4.36 @@ -34,81 +31,590 @@
    4.37    vor_minus_mal         "- a * b = (-a) * b"
    4.38  
    4.39    (*commute with invariant (a.b).c -association*)
    4.40 -  tausche_plus		"[| b ist_monom; a kleiner b  |] ==> \
    4.41 -			\(b + a) = (a + b)"
    4.42 -  tausche_minus		"[| b ist_monom; a kleiner b  |] ==> \
    4.43 -			\(b - a) = (-a + b)"
    4.44 -  tausche_vor_plus	"[| b ist_monom; a kleiner b  |] ==> \
    4.45 -			\(- b + a) = (a - b)"
    4.46 -  tausche_vor_minus	"[| b ist_monom; a kleiner b  |] ==> \
    4.47 -			\(- b - a) = (-a - b)"
    4.48 +  tausche_plus		"[| b ist_monom; a kleiner b  |] ==> 
    4.49 +			 (b + a) = (a + b)"
    4.50 +  tausche_minus		"[| b ist_monom; a kleiner b  |] ==> 
    4.51 +			 (b - a) = (-a + b)"
    4.52 +  tausche_vor_plus	"[| b ist_monom; a kleiner b  |] ==> 
    4.53 +			 (- b + a) = (a - b)"
    4.54 +  tausche_vor_minus	"[| b ist_monom; a kleiner b  |] ==> 
    4.55 +			 (- b - a) = (-a - b)"
    4.56    tausche_plus_plus	"b kleiner c ==> (a + c + b) = (a + b + c)"
    4.57    tausche_plus_minus	"b kleiner c ==> (a + c - b) = (a - b + c)"
    4.58    tausche_minus_plus	"b kleiner c ==> (a - c + b) = (a + b - c)"
    4.59    tausche_minus_minus	"b kleiner c ==> (a - c - b) = (a - b - c)"
    4.60  
    4.61    (*commute with invariant (a.b).c -association*)
    4.62 -  tausche_mal		"[| b is_atom; a kleiner b  |] ==> \
    4.63 -			\(b * a) = (a * b)"
    4.64 -  tausche_vor_mal	"[| b is_atom; a kleiner b  |] ==> \
    4.65 -			\(-b * a) = (-a * b)"
    4.66 -  tausche_mal_mal	"[| c is_atom; b kleiner c  |] ==> \
    4.67 -			\(x * c * b) = (x * b * c)"
    4.68 +  tausche_mal		"[| b is_atom; a kleiner b  |] ==> 
    4.69 +			 (b * a) = (a * b)"
    4.70 +  tausche_vor_mal	"[| b is_atom; a kleiner b  |] ==> 
    4.71 +			 (-b * a) = (-a * b)"
    4.72 +  tausche_mal_mal	"[| c is_atom; b kleiner c  |] ==> 
    4.73 +			 (x * c * b) = (x * b * c)"
    4.74    x_quadrat             "(x * a) * a = x * a ^^^ 2"
    4.75  
    4.76  
    4.77 -  subtrahiere               "[| l is_const; m is_const |] ==>  \
    4.78 -			    \m * v - l * v = (m - l) * v"
    4.79 -  subtrahiere_von_1         "[| l is_const |] ==>  \
    4.80 -			    \v - l * v = (1 - l) * v"
    4.81 -  subtrahiere_1             "[| l is_const; m is_const |] ==>  \
    4.82 -			    \m * v - v = (m - 1) * v"
    4.83 +  subtrahiere               "[| l is_const; m is_const |] ==>  
    4.84 +			     m * v - l * v = (m - l) * v"
    4.85 +  subtrahiere_von_1         "[| l is_const |] ==>  
    4.86 +			     v - l * v = (1 - l) * v"
    4.87 +  subtrahiere_1             "[| l is_const; m is_const |] ==>  
    4.88 +			     m * v - v = (m - 1) * v"
    4.89  
    4.90 -  subtrahiere_x_plus_minus  "[| l is_const; m is_const |] ==>  \
    4.91 -			    \(x + m * v) - l * v = x + (m - l) * v"
    4.92 -  subtrahiere_x_plus1_minus "[| l is_const |] ==>  \
    4.93 -			    \(x + v) - l * v = x + (1 - l) * v"
    4.94 -  subtrahiere_x_plus_minus1 "[| m is_const |] ==>  \
    4.95 -			    \(x + m * v) - v = x + (m - 1) * v"
    4.96 +  subtrahiere_x_plus_minus  "[| l is_const; m is_const |] ==>  
    4.97 +			     (x + m * v) - l * v = x + (m - l) * v"
    4.98 +  subtrahiere_x_plus1_minus "[| l is_const |] ==>  
    4.99 +			     (x + v) - l * v = x + (1 - l) * v"
   4.100 +  subtrahiere_x_plus_minus1 "[| m is_const |] ==>  
   4.101 +			     (x + m * v) - v = x + (m - 1) * v"
   4.102  
   4.103 -  subtrahiere_x_minus_plus  "[| l is_const; m is_const |] ==>  \
   4.104 -			    \(x - m * v) + l * v = x + (-m + l) * v"
   4.105 -  subtrahiere_x_minus1_plus "[| l is_const |] ==>  \
   4.106 -			    \(x - v) + l * v = x + (-1 + l) * v"
   4.107 -  subtrahiere_x_minus_plus1 "[| m is_const |] ==>  \
   4.108 -			    \(x - m * v) + v = x + (-m + 1) * v"
   4.109 +  subtrahiere_x_minus_plus  "[| l is_const; m is_const |] ==>  
   4.110 +			     (x - m * v) + l * v = x + (-m + l) * v"
   4.111 +  subtrahiere_x_minus1_plus "[| l is_const |] ==>  
   4.112 +			     (x - v) + l * v = x + (-1 + l) * v"
   4.113 +  subtrahiere_x_minus_plus1 "[| m is_const |] ==>  
   4.114 +			     (x - m * v) + v = x + (-m + 1) * v"
   4.115  
   4.116 -  subtrahiere_x_minus_minus "[| l is_const; m is_const |] ==>  \
   4.117 -			    \(x - m * v) - l * v = x + (-m - l) * v"
   4.118 -  subtrahiere_x_minus1_minus"[| l is_const |] ==>  \
   4.119 -			    \(x - v) - l * v = x + (-1 - l) * v"
   4.120 -  subtrahiere_x_minus_minus1"[| m is_const |] ==>  \
   4.121 -			    \(x - m * v) - v = x + (-m - 1) * v"
   4.122 +  subtrahiere_x_minus_minus "[| l is_const; m is_const |] ==>  
   4.123 +			     (x - m * v) - l * v = x + (-m - l) * v"
   4.124 +  subtrahiere_x_minus1_minus"[| l is_const |] ==>  
   4.125 +			     (x - v) - l * v = x + (-1 - l) * v"
   4.126 +  subtrahiere_x_minus_minus1"[| m is_const |] ==>  
   4.127 +			     (x - m * v) - v = x + (-m - 1) * v"
   4.128  
   4.129  
   4.130 -  addiere_vor_minus         "[| l is_const; m is_const |] ==>  \
   4.131 -			    \- (l * v) +  m * v = (-l + m) * v"
   4.132 -  addiere_eins_vor_minus    "[| m is_const |] ==>  \
   4.133 -			    \-  v +  m * v = (-1 + m) * v"
   4.134 -  subtrahiere_vor_minus     "[| l is_const; m is_const |] ==>  \
   4.135 -			    \- (l * v) -  m * v = (-l - m) * v"
   4.136 -  subtrahiere_eins_vor_minus"[| m is_const |] ==>  \
   4.137 -			    \-  v -  m * v = (-1 - m) * v"
   4.138 +  addiere_vor_minus         "[| l is_const; m is_const |] ==>  
   4.139 +			     - (l * v) +  m * v = (-l + m) * v"
   4.140 +  addiere_eins_vor_minus    "[| m is_const |] ==>  
   4.141 +			     -  v +  m * v = (-1 + m) * v"
   4.142 +  subtrahiere_vor_minus     "[| l is_const; m is_const |] ==>  
   4.143 +			     - (l * v) -  m * v = (-l - m) * v"
   4.144 +  subtrahiere_eins_vor_minus"[| m is_const |] ==>  
   4.145 +			     -  v -  m * v = (-1 - m) * v"
   4.146  
   4.147 -  vorzeichen_minus_weg1  "l kleiner 0 ==> a + l * b = a - -1*l * b"
   4.148 -  vorzeichen_minus_weg2  "l kleiner 0 ==> a - l * b = a + -1*l * b"
   4.149 -  vorzeichen_minus_weg3  "l kleiner 0 ==> k + a - l * b = k + a + -1*l * b"
   4.150 -  vorzeichen_minus_weg4  "l kleiner 0 ==> k - a - l * b = k - a + -1*l * b"
   4.151 +  vorzeichen_minus_weg1      "l kleiner 0 ==> a + l * b = a - -1*l * b"
   4.152 +  vorzeichen_minus_weg2      "l kleiner 0 ==> a - l * b = a + -1*l * b"
   4.153 +  vorzeichen_minus_weg3      "l kleiner 0 ==> k + a - l * b = k + a + -1*l * b"
   4.154 +  vorzeichen_minus_weg4      "l kleiner 0 ==> k - a - l * b = k - a + -1*l * b"
   4.155  
   4.156    (*klammer_plus_plus = (real_add_assoc RS sym)*)
   4.157 -  klammer_plus_minus     "a + (b - c) = (a + b) - c"
   4.158 -  klammer_minus_plus     "a - (b + c) = (a - b) - c"
   4.159 -  klammer_minus_minus    "a - (b - c) = (a - b) + c"
   4.160 +  klammer_plus_minus          "a + (b - c) = (a + b) - c"
   4.161 +  klammer_minus_plus          "a - (b + c) = (a - b) - c"
   4.162 +  klammer_minus_minus         "a - (b - c) = (a - b) + c"
   4.163  
   4.164 -  klammer_mult_minus      "a * (b - c) = a * b - a * c"
   4.165 -  klammer_minus_mult      "(b - c) * a = b * a - c * a"
   4.166 +  klammer_mult_minus          "a * (b - c) = a * b - a * c"
   4.167 +  klammer_minus_mult          "(b - c) * a = b * a - c * a"
   4.168  
   4.169 +ML {*
   4.170 +(** eval functions **)
   4.171  
   4.172 +(*. get the identifier from specific monomials; see fun ist_monom .*)
   4.173 +(*HACK.WN080107*)
   4.174 +fun increase str = 
   4.175 +    let val s::ss = explode str
   4.176 +    in implode ((chr (ord s + 1))::ss) end;
   4.177 +fun identifier (Free (id,_)) = id                            (* 2,   a   *)
   4.178 +  | identifier (Const ("op *", _) $ Free (num, _) $ Free (id, _)) = 
   4.179 +    id                                                       (* 2*a, a*b *)
   4.180 +  | identifier (Const ("op *", _) $                          (* 3*a*b    *)
   4.181 +		     (Const ("op *", _) $
   4.182 +			    Free (num, _) $ Free _) $ Free (id, _)) = 
   4.183 +    if is_numeral num then id
   4.184 +    else "|||||||||||||"
   4.185 +  | identifier (Const ("Atools.pow", _) $ Free (base, _) $ Free (exp, _)) =
   4.186 +    if is_numeral base then "|||||||||||||"                  (* a^2      *)
   4.187 +    else (*increase*) base
   4.188 +  | identifier (Const ("op *", _) $ Free (num, _) $          (* 3*a^2    *)
   4.189 +		     (Const ("Atools.pow", _) $
   4.190 +			    Free (base, _) $ Free (exp, _))) = 
   4.191 +    if is_numeral num andalso not (is_numeral base) then (*increase*) base
   4.192 +    else "|||||||||||||"
   4.193 +  | identifier _ = "|||||||||||||"(*the "largest" string*);
   4.194 +
   4.195 +(*("kleiner", ("PolyMinus.kleiner", eval_kleiner ""))*)
   4.196 +(* order "by alphabet" w.r.t. var: num < (var | num*var) > (var*var | ..) *)
   4.197 +fun eval_kleiner _ _ (p as (Const ("PolyMinus.kleiner",_) $ a $ b)) _  =
   4.198 +     if is_num b then
   4.199 +	 if is_num a then (*123 kleiner 32 = True !!!*)
   4.200 +	     if int_of_Free a < int_of_Free b then 
   4.201 +		 SOME ((term2str p) ^ " = True",
   4.202 +		       Trueprop $ (mk_equality (p, HOLogic.true_const)))
   4.203 +	     else SOME ((term2str p) ^ " = False",
   4.204 +			Trueprop $ (mk_equality (p, HOLogic.false_const)))
   4.205 +	 else (* -1 * -2 kleiner 0 *)
   4.206 +	     SOME ((term2str p) ^ " = False",
   4.207 +		   Trueprop $ (mk_equality (p, HOLogic.false_const)))
   4.208 +    else
   4.209 +	if identifier a < identifier b then 
   4.210 +	     SOME ((term2str p) ^ " = True",
   4.211 +		  Trueprop $ (mk_equality (p, HOLogic.true_const)))
   4.212 +	else SOME ((term2str p) ^ " = False",
   4.213 +		   Trueprop $ (mk_equality (p, HOLogic.false_const)))
   4.214 +  | eval_kleiner _ _ _ _ =  NONE;
   4.215 +
   4.216 +fun ist_monom (Free (id,_)) = true
   4.217 +  | ist_monom (Const ("op *", _) $ Free (num, _) $ Free (id, _)) = 
   4.218 +    if is_numeral num then true else false
   4.219 +  | ist_monom _ = false;
   4.220 +(*. this function only accepts the most simple monoms       vvvvvvvvvv .*)
   4.221 +fun ist_monom (Free (id,_)) = true                          (* 2,   a   *)
   4.222 +  | ist_monom (Const ("op *", _) $ Free _ $ Free (id, _)) = (* 2*a, a*b *)
   4.223 +    if is_numeral id then false else true
   4.224 +  | ist_monom (Const ("op *", _) $                          (* 3*a*b    *)
   4.225 +		     (Const ("op *", _) $
   4.226 +			    Free (num, _) $ Free _) $ Free (id, _)) =
   4.227 +    if is_numeral num andalso not (is_numeral id) then true else false
   4.228 +  | ist_monom (Const ("Atools.pow", _) $ Free (base, _) $ Free (exp, _)) = 
   4.229 +    true                                                    (* a^2      *)
   4.230 +  | ist_monom (Const ("op *", _) $ Free (num, _) $          (* 3*a^2    *)
   4.231 +		     (Const ("Atools.pow", _) $
   4.232 +			    Free (base, _) $ Free (exp, _))) = 
   4.233 +    if is_numeral num then true else false
   4.234 +  | ist_monom _ = false;
   4.235 +
   4.236 +(* is this a univariate monomial ? *)
   4.237 +(*("ist_monom", ("PolyMinus.ist'_monom", eval_ist_monom ""))*)
   4.238 +fun eval_ist_monom _ _ (p as (Const ("PolyMinus.ist'_monom",_) $ a)) _  =
   4.239 +    if ist_monom a  then 
   4.240 +	SOME ((term2str p) ^ " = True",
   4.241 +	      Trueprop $ (mk_equality (p, HOLogic.true_const)))
   4.242 +    else SOME ((term2str p) ^ " = False",
   4.243 +	       Trueprop $ (mk_equality (p, HOLogic.false_const)))
   4.244 +  | eval_ist_monom _ _ _ _ =  NONE;
   4.245 +
   4.246 +
   4.247 +(** rewrite order **)
   4.248 +
   4.249 +(** rulesets **)
   4.250 +
   4.251 +val erls_ordne_alphabetisch =
   4.252 +    append_rls "erls_ordne_alphabetisch" e_rls
   4.253 +	       [Calc ("PolyMinus.kleiner", eval_kleiner ""),
   4.254 +		Calc ("PolyMinus.ist'_monom", eval_ist_monom "")
   4.255 +		];
   4.256 +
   4.257 +val ordne_alphabetisch = 
   4.258 +  Rls{id = "ordne_alphabetisch", preconds = [], 
   4.259 +      rew_ord = ("dummy_ord", dummy_ord), srls = Erls, calc = [],
   4.260 +      erls = erls_ordne_alphabetisch, 
   4.261 +      rules = [Thm ("tausche_plus",num_str tausche_plus),
   4.262 +	       (*"b kleiner a ==> (b + a) = (a + b)"*)
   4.263 +	       Thm ("tausche_minus",num_str tausche_minus),
   4.264 +	       (*"b kleiner a ==> (b - a) = (-a + b)"*)
   4.265 +	       Thm ("tausche_vor_plus",num_str tausche_vor_plus),
   4.266 +	       (*"[| b ist_monom; a kleiner b  |] ==> (- b + a) = (a - b)"*)
   4.267 +	       Thm ("tausche_vor_minus",num_str tausche_vor_minus),
   4.268 +	       (*"[| b ist_monom; a kleiner b  |] ==> (- b - a) = (-a - b)"*)
   4.269 +	       Thm ("tausche_plus_plus",num_str tausche_plus_plus),
   4.270 +	       (*"c kleiner b ==> (a + c + b) = (a + b + c)"*)
   4.271 +	       Thm ("tausche_plus_minus",num_str tausche_plus_minus),
   4.272 +	       (*"c kleiner b ==> (a + c - b) = (a - b + c)"*)
   4.273 +	       Thm ("tausche_minus_plus",num_str tausche_minus_plus),
   4.274 +	       (*"c kleiner b ==> (a - c + b) = (a + b - c)"*)
   4.275 +	       Thm ("tausche_minus_minus",num_str tausche_minus_minus)
   4.276 +	       (*"c kleiner b ==> (a - c - b) = (a - b - c)"*)
   4.277 +	       ], scr = EmptyScr}:rls;
   4.278 +
   4.279 +val fasse_zusammen = 
   4.280 +    Rls{id = "fasse_zusammen", preconds = [], 
   4.281 +	rew_ord = ("dummy_ord", dummy_ord),
   4.282 +	erls = append_rls "erls_fasse_zusammen" e_rls 
   4.283 +			  [Calc ("Atools.is'_const",eval_const "#is_const_")], 
   4.284 +	srls = Erls, calc = [],
   4.285 +	rules = 
   4.286 +	[Thm ("real_num_collect",num_str real_num_collect), 
   4.287 +	 (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
   4.288 +	 Thm ("real_num_collect_assoc_r",num_str real_num_collect_assoc_r),
   4.289 +	 (*"[| l is_const; m..|] ==>  (k + m * n) + l * n = k + (l + m)*n"*)
   4.290 +	 Thm ("real_one_collect",num_str real_one_collect),	
   4.291 +	 (*"m is_const ==> n + m * n = (1 + m) * n"*)
   4.292 +	 Thm ("real_one_collect_assoc_r",num_str real_one_collect_assoc_r), 
   4.293 +	 (*"m is_const ==> (k + n) + m * n = k + (m + 1) * n"*)
   4.294 +
   4.295 +
   4.296 +	 Thm ("subtrahiere",num_str subtrahiere),
   4.297 +	 (*"[| l is_const; m is_const |] ==> m * v - l * v = (m - l) * v"*)
   4.298 +	 Thm ("subtrahiere_von_1",num_str subtrahiere_von_1),
   4.299 +	 (*"[| l is_const |] ==> v - l * v = (1 - l) * v"*)
   4.300 +	 Thm ("subtrahiere_1",num_str subtrahiere_1),
   4.301 +	 (*"[| l is_const; m is_const |] ==> m * v - v = (m - 1) * v"*)
   4.302 +
   4.303 +	 Thm ("subtrahiere_x_plus_minus",num_str subtrahiere_x_plus_minus), 
   4.304 +	 (*"[| l is_const; m..|] ==> (k + m * n) - l * n = k + ( m - l) * n"*)
   4.305 +	 Thm ("subtrahiere_x_plus1_minus",num_str subtrahiere_x_plus1_minus),
   4.306 +	 (*"[| l is_const |] ==> (x + v) - l * v = x + (1 - l) * v"*)
   4.307 +	 Thm ("subtrahiere_x_plus_minus1",num_str subtrahiere_x_plus_minus1),
   4.308 +	 (*"[| m is_const |] ==> (x + m * v) - v = x + (m - 1) * v"*)
   4.309 +
   4.310 +	 Thm ("subtrahiere_x_minus_plus",num_str subtrahiere_x_minus_plus), 
   4.311 +	 (*"[| l is_const; m..|] ==> (k - m * n) + l * n = k + (-m + l) * n"*)
   4.312 +	 Thm ("subtrahiere_x_minus1_plus",num_str subtrahiere_x_minus1_plus),
   4.313 +	 (*"[| l is_const |] ==> (x - v) + l * v = x + (-1 + l) * v"*)
   4.314 +	 Thm ("subtrahiere_x_minus_plus1",num_str subtrahiere_x_minus_plus1),
   4.315 +	 (*"[| m is_const |] ==> (x - m * v) + v = x + (-m + 1) * v"*)
   4.316 +
   4.317 +	 Thm ("subtrahiere_x_minus_minus",num_str subtrahiere_x_minus_minus), 
   4.318 +	 (*"[| l is_const; m..|] ==> (k - m * n) - l * n = k + (-m - l) * n"*)
   4.319 +	 Thm ("subtrahiere_x_minus1_minus",num_str subtrahiere_x_minus1_minus),
   4.320 +	 (*"[| l is_const |] ==> (x - v) - l * v = x + (-1 - l) * v"*)
   4.321 +	 Thm ("subtrahiere_x_minus_minus1",num_str subtrahiere_x_minus_minus1),
   4.322 +	 (*"[| m is_const |] ==> (x - m * v) - v = x + (-m - 1) * v"*)
   4.323 +	 
   4.324 +	 Calc ("op +", eval_binop "#add_"),
   4.325 +	 Calc ("op -", eval_binop "#subtr_"),
   4.326 +	 
   4.327 +	 (*MG: Reihenfolge der folgenden 2 Thm muss so bleiben, wegen
   4.328 +           (a+a)+a --> a + 2*a --> 3*a and not (a+a)+a --> 2*a + a *)
   4.329 +	 Thm ("real_mult_2_assoc_r",num_str real_mult_2_assoc_r),
   4.330 +	 (*"(k + z1) + z1 = k + 2 * z1"*)
   4.331 +	 Thm ("sym_real_mult_2",num_str (real_mult_2 RS sym)),
   4.332 +	 (*"z1 + z1 = 2 * z1"*)
   4.333 +
   4.334 +	 Thm ("addiere_vor_minus",num_str addiere_vor_minus),
   4.335 +	 (*"[| l is_const; m is_const |] ==> -(l * v) +  m * v = (-l + m) *v"*)
   4.336 +	 Thm ("addiere_eins_vor_minus",num_str addiere_eins_vor_minus),
   4.337 +	 (*"[| m is_const |] ==> -  v +  m * v = (-1 + m) * v"*)
   4.338 +	 Thm ("subtrahiere_vor_minus",num_str subtrahiere_vor_minus),
   4.339 +	 (*"[| l is_const; m is_const |] ==> -(l * v) -  m * v = (-l - m) *v"*)
   4.340 +	 Thm ("subtrahiere_eins_vor_minus",num_str subtrahiere_eins_vor_minus)
   4.341 +	 (*"[| m is_const |] ==> -  v -  m * v = (-1 - m) * v"*)
   4.342 +	 
   4.343 +	 ], scr = EmptyScr}:rls;
   4.344 +    
   4.345 +val verschoenere = 
   4.346 +  Rls{id = "verschoenere", preconds = [], 
   4.347 +      rew_ord = ("dummy_ord", dummy_ord), srls = Erls, calc = [],
   4.348 +      erls = append_rls "erls_verschoenere" e_rls 
   4.349 +			[Calc ("PolyMinus.kleiner", eval_kleiner "")], 
   4.350 +      rules = [Thm ("vorzeichen_minus_weg1",num_str vorzeichen_minus_weg1),
   4.351 +	       (*"l kleiner 0 ==> a + l * b = a - -l * b"*)
   4.352 +	       Thm ("vorzeichen_minus_weg2",num_str vorzeichen_minus_weg2),
   4.353 +	       (*"l kleiner 0 ==> a - l * b = a + -l * b"*)
   4.354 +	       Thm ("vorzeichen_minus_weg3",num_str vorzeichen_minus_weg3),
   4.355 +	       (*"l kleiner 0 ==> k + a - l * b = k + a + -l * b"*)
   4.356 +	       Thm ("vorzeichen_minus_weg4",num_str vorzeichen_minus_weg4),
   4.357 +	       (*"l kleiner 0 ==> k - a - l * b = k - a + -l * b"*)
   4.358 +
   4.359 +	       Calc ("op *", eval_binop "#mult_"),
   4.360 +
   4.361 +	       Thm ("real_mult_0",num_str real_mult_0),    
   4.362 +	       (*"0 * z = 0"*)
   4.363 +	       Thm ("real_mult_1",num_str real_mult_1),     
   4.364 +	       (*"1 * z = z"*)
   4.365 +	       Thm ("real_add_zero_left",num_str real_add_zero_left),
   4.366 +	       (*"0 + z = z"*)
   4.367 +	       Thm ("null_minus",num_str null_minus),
   4.368 +	       (*"0 - a = -a"*)
   4.369 +	       Thm ("vor_minus_mal",num_str vor_minus_mal)
   4.370 +	       (*"- a * b = (-a) * b"*)
   4.371 +
   4.372 +	       (*Thm ("",num_str ),*)
   4.373 +	       (**)
   4.374 +	       ], scr = EmptyScr}:rls (*end verschoenere*);
   4.375 +
   4.376 +val klammern_aufloesen = 
   4.377 +  Rls{id = "klammern_aufloesen", preconds = [], 
   4.378 +      rew_ord = ("dummy_ord", dummy_ord), srls = Erls, calc = [], erls = Erls, 
   4.379 +      rules = [Thm ("sym_real_add_assoc",num_str (real_add_assoc RS sym)),
   4.380 +	       (*"a + (b + c) = (a + b) + c"*)
   4.381 +	       Thm ("klammer_plus_minus",num_str klammer_plus_minus),
   4.382 +	       (*"a + (b - c) = (a + b) - c"*)
   4.383 +	       Thm ("klammer_minus_plus",num_str klammer_minus_plus),
   4.384 +	       (*"a - (b + c) = (a - b) - c"*)
   4.385 +	       Thm ("klammer_minus_minus",num_str klammer_minus_minus)
   4.386 +	       (*"a - (b - c) = (a - b) + c"*)
   4.387 +	       ], scr = EmptyScr}:rls;
   4.388 +
   4.389 +val klammern_ausmultiplizieren = 
   4.390 +  Rls{id = "klammern_ausmultiplizieren", preconds = [], 
   4.391 +      rew_ord = ("dummy_ord", dummy_ord), srls = Erls, calc = [], erls = Erls, 
   4.392 +      rules = [Thm ("real_add_mult_distrib" ,num_str real_add_mult_distrib),
   4.393 +	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   4.394 +	       Thm ("real_add_mult_distrib2",num_str real_add_mult_distrib2),
   4.395 +	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   4.396 +	       
   4.397 +	       Thm ("klammer_mult_minus",num_str klammer_mult_minus),
   4.398 +	       (*"a * (b - c) = a * b - a * c"*)
   4.399 +	       Thm ("klammer_minus_mult",num_str klammer_minus_mult)
   4.400 +	       (*"(b - c) * a = b * a - c * a"*)
   4.401 +
   4.402 +	       (*Thm ("",num_str ),
   4.403 +	       (*""*)*)
   4.404 +	       ], scr = EmptyScr}:rls;
   4.405 +
   4.406 +val ordne_monome = 
   4.407 +  Rls{id = "ordne_monome", preconds = [], 
   4.408 +      rew_ord = ("dummy_ord", dummy_ord), srls = Erls, calc = [], 
   4.409 +      erls = append_rls "erls_ordne_monome" e_rls
   4.410 +	       [Calc ("PolyMinus.kleiner", eval_kleiner ""),
   4.411 +		Calc ("Atools.is'_atom", eval_is_atom "")
   4.412 +		], 
   4.413 +      rules = [Thm ("tausche_mal",num_str tausche_mal),
   4.414 +	       (*"[| b is_atom; a kleiner b  |] ==> (b * a) = (a * b)"*)
   4.415 +	       Thm ("tausche_vor_mal",num_str tausche_vor_mal),
   4.416 +	       (*"[| b is_atom; a kleiner b  |] ==> (-b * a) = (-a * b)"*)
   4.417 +	       Thm ("tausche_mal_mal",num_str tausche_mal_mal),
   4.418 +	       (*"[| c is_atom; b kleiner c  |] ==> (a * c * b) = (a * b *c)"*)
   4.419 +	       Thm ("x_quadrat",num_str x_quadrat)
   4.420 +	       (*"(x * a) * a = x * a ^^^ 2"*)
   4.421 +
   4.422 +	       (*Thm ("",num_str ),
   4.423 +	       (*""*)*)
   4.424 +	       ], scr = EmptyScr}:rls;
   4.425 +
   4.426 +
   4.427 +val rls_p_33 = 
   4.428 +    append_rls "rls_p_33" e_rls
   4.429 +	       [Rls_ ordne_alphabetisch,
   4.430 +		Rls_ fasse_zusammen,
   4.431 +		Rls_ verschoenere
   4.432 +		];
   4.433 +val rls_p_34 = 
   4.434 +    append_rls "rls_p_34" e_rls
   4.435 +	       [Rls_ klammern_aufloesen,
   4.436 +		Rls_ ordne_alphabetisch,
   4.437 +		Rls_ fasse_zusammen,
   4.438 +		Rls_ verschoenere
   4.439 +		];
   4.440 +val rechnen = 
   4.441 +    append_rls "rechnen" e_rls
   4.442 +	       [Calc ("op *", eval_binop "#mult_"),
   4.443 +		Calc ("op +", eval_binop "#add_"),
   4.444 +		Calc ("op -", eval_binop "#subtr_")
   4.445 +		];
   4.446 +
   4.447 +ruleset' := 
   4.448 +overwritelthy thy (!ruleset',
   4.449 +		   [("ordne_alphabetisch", prep_rls ordne_alphabetisch),
   4.450 +		    ("fasse_zusammen", prep_rls fasse_zusammen),
   4.451 +		    ("verschoenere", prep_rls verschoenere),
   4.452 +		    ("ordne_monome", prep_rls ordne_monome),
   4.453 +		    ("klammern_aufloesen", prep_rls klammern_aufloesen),
   4.454 +		    ("klammern_ausmultiplizieren", 
   4.455 +		     prep_rls klammern_ausmultiplizieren)
   4.456 +		    ]);
   4.457 +
   4.458 +(** problems **)
   4.459 +
   4.460 +store_pbt
   4.461 + (prep_pbt PolyMinus.thy "pbl_vereinf_poly" [] e_pblID
   4.462 + (["polynom","vereinfachen"],
   4.463 +  [], Erls, NONE, []));
   4.464 +
   4.465 +store_pbt
   4.466 + (prep_pbt PolyMinus.thy "pbl_vereinf_poly_minus" [] e_pblID
   4.467 + (["plus_minus","polynom","vereinfachen"],
   4.468 +  [("#Given" ,["term t_"]),
   4.469 +   ("#Where" ,["t_ is_polyexp",
   4.470 +	       "Not (matchsub (?a + (?b + ?c)) t_ | " ^
   4.471 +	       "     matchsub (?a + (?b - ?c)) t_ | " ^
   4.472 +	       "     matchsub (?a - (?b + ?c)) t_ | " ^
   4.473 +	       "     matchsub (?a + (?b - ?c)) t_ )",
   4.474 +	       "Not (matchsub (?a * (?b + ?c)) t_ | " ^
   4.475 +	       "     matchsub (?a * (?b - ?c)) t_ | " ^
   4.476 +	       "     matchsub ((?b + ?c) * ?a) t_ | " ^
   4.477 +	       "     matchsub ((?b - ?c) * ?a) t_ )"]),
   4.478 +   ("#Find"  ,["normalform n_"])
   4.479 +  ],
   4.480 +  append_rls "prls_pbl_vereinf_poly" e_rls 
   4.481 +	     [Calc ("Poly.is'_polyexp", eval_is_polyexp ""),
   4.482 +	      Calc ("Tools.matchsub", eval_matchsub ""),
   4.483 +	      Thm ("or_true",or_true),
   4.484 +	      (*"(?a | True) = True"*)
   4.485 +	      Thm ("or_false",or_false),
   4.486 +	      (*"(?a | False) = ?a"*)
   4.487 +	      Thm ("not_true",num_str not_true),
   4.488 +	      (*"(~ True) = False"*)
   4.489 +	      Thm ("not_false",num_str not_false)
   4.490 +	      (*"(~ False) = True"*)], 
   4.491 +  SOME "Vereinfache t_", 
   4.492 +  [["simplification","for_polynomials","with_minus"]]));
   4.493 +
   4.494 +store_pbt
   4.495 + (prep_pbt PolyMinus.thy "pbl_vereinf_poly_klammer" [] e_pblID
   4.496 + (["klammer","polynom","vereinfachen"],
   4.497 +  [("#Given" ,["term t_"]),
   4.498 +   ("#Where" ,["t_ is_polyexp",
   4.499 +	       "Not (matchsub (?a * (?b + ?c)) t_ | " ^
   4.500 +	       "     matchsub (?a * (?b - ?c)) t_ | " ^
   4.501 +	       "     matchsub ((?b + ?c) * ?a) t_ | " ^
   4.502 +	       "     matchsub ((?b - ?c) * ?a) t_ )"]),
   4.503 +   ("#Find"  ,["normalform n_"])
   4.504 +  ],
   4.505 +  append_rls "prls_pbl_vereinf_poly_klammer" e_rls [Calc ("Poly.is'_polyexp", eval_is_polyexp ""),
   4.506 +	      Calc ("Tools.matchsub", eval_matchsub ""),
   4.507 +	      Thm ("or_true",or_true),
   4.508 +	      (*"(?a | True) = True"*)
   4.509 +	      Thm ("or_false",or_false),
   4.510 +	      (*"(?a | False) = ?a"*)
   4.511 +	      Thm ("not_true",num_str not_true),
   4.512 +	      (*"(~ True) = False"*)
   4.513 +	      Thm ("not_false",num_str not_false)
   4.514 +	      (*"(~ False) = True"*)], 
   4.515 +  SOME "Vereinfache t_", 
   4.516 +  [["simplification","for_polynomials","with_parentheses"]]));
   4.517 +
   4.518 +store_pbt
   4.519 + (prep_pbt PolyMinus.thy "pbl_vereinf_poly_klammer_mal" [] e_pblID
   4.520 + (["binom_klammer","polynom","vereinfachen"],
   4.521 +  [("#Given" ,["term t_"]),
   4.522 +   ("#Where" ,["t_ is_polyexp"]),
   4.523 +   ("#Find"  ,["normalform n_"])
   4.524 +  ],
   4.525 +  append_rls "e_rls" e_rls [(*for preds in where_*)
   4.526 +			    Calc ("Poly.is'_polyexp", eval_is_polyexp "")], 
   4.527 +  SOME "Vereinfache t_", 
   4.528 +  [["simplification","for_polynomials","with_parentheses_mult"]]));
   4.529 +
   4.530 +store_pbt
   4.531 + (prep_pbt PolyMinus.thy "pbl_probe" [] e_pblID
   4.532 + (["probe"],
   4.533 +  [], Erls, NONE, []));
   4.534 +
   4.535 +store_pbt
   4.536 + (prep_pbt PolyMinus.thy "pbl_probe_poly" [] e_pblID
   4.537 + (["polynom","probe"],
   4.538 +  [("#Given" ,["Pruefe e_", "mitWert ws_"]),
   4.539 +   ("#Where" ,["e_ is_polyexp"]),
   4.540 +   ("#Find"  ,["Geprueft p_"])
   4.541 +  ],
   4.542 +  append_rls "prls_pbl_probe_poly" 
   4.543 +	     e_rls [(*for preds in where_*)
   4.544 +		    Calc ("Poly.is'_polyexp", eval_is_polyexp "")], 
   4.545 +  SOME "Probe e_ ws_", 
   4.546 +  [["probe","fuer_polynom"]]));
   4.547 +
   4.548 +store_pbt
   4.549 + (prep_pbt PolyMinus.thy "pbl_probe_bruch" [] e_pblID
   4.550 + (["bruch","probe"],
   4.551 +  [("#Given" ,["Pruefe e_", "mitWert ws_"]),
   4.552 +   ("#Where" ,["e_ is_ratpolyexp"]),
   4.553 +   ("#Find"  ,["Geprueft p_"])
   4.554 +  ],
   4.555 +  append_rls "prls_pbl_probe_bruch"
   4.556 +	     e_rls [(*for preds in where_*)
   4.557 +		    Calc ("Rational.is'_ratpolyexp", eval_is_ratpolyexp "")], 
   4.558 +  SOME "Probe e_ ws_", 
   4.559 +  [["probe","fuer_bruch"]]));
   4.560 +
   4.561 +
   4.562 +(** methods **)
   4.563 +
   4.564 +store_met
   4.565 +    (prep_met PolyMinus.thy "met_simp_poly_minus" [] e_metID
   4.566 +	      (["simplification","for_polynomials","with_minus"],
   4.567 +	       [("#Given" ,["term t_"]),
   4.568 +		("#Where" ,["t_ is_polyexp",
   4.569 +	       "Not (matchsub (?a + (?b + ?c)) t_ | " ^
   4.570 +	       "     matchsub (?a + (?b - ?c)) t_ | " ^
   4.571 +	       "     matchsub (?a - (?b + ?c)) t_ | " ^
   4.572 +	       "     matchsub (?a + (?b - ?c)) t_ )"]),
   4.573 +		("#Find"  ,["normalform n_"])
   4.574 +		],
   4.575 +	       {rew_ord'="tless_true", rls' = e_rls, calc = [], srls = e_rls, 
   4.576 +		prls = append_rls "prls_met_simp_poly_minus" e_rls 
   4.577 +				  [Calc ("Poly.is'_polyexp", eval_is_polyexp ""),
   4.578 +	      Calc ("Tools.matchsub", eval_matchsub ""),
   4.579 +	      Thm ("and_true",and_true),
   4.580 +	      (*"(?a & True) = ?a"*)
   4.581 +	      Thm ("and_false",and_false),
   4.582 +	      (*"(?a & False) = False"*)
   4.583 +	      Thm ("not_true",num_str not_true),
   4.584 +	      (*"(~ True) = False"*)
   4.585 +	      Thm ("not_false",num_str not_false)
   4.586 +	      (*"(~ False) = True"*)],
   4.587 +		crls = e_rls, nrls = rls_p_33},
   4.588 +"Script SimplifyScript (t_::real) =                   " ^
   4.589 +"  ((Repeat((Try (Rewrite_Set ordne_alphabetisch False)) @@  " ^
   4.590 +"           (Try (Rewrite_Set fasse_zusammen     False)) @@  " ^
   4.591 +"           (Try (Rewrite_Set verschoenere       False)))) t_)"
   4.592 +	       ));
   4.593 +
   4.594 +store_met
   4.595 +    (prep_met PolyMinus.thy "met_simp_poly_parenth" [] e_metID
   4.596 +	      (["simplification","for_polynomials","with_parentheses"],
   4.597 +	       [("#Given" ,["term t_"]),
   4.598 +		("#Where" ,["t_ is_polyexp"]),
   4.599 +		("#Find"  ,["normalform n_"])
   4.600 +		],
   4.601 +	       {rew_ord'="tless_true", rls' = e_rls, calc = [], srls = e_rls, 
   4.602 +		prls = append_rls "simplification_for_polynomials_prls" e_rls 
   4.603 +				  [(*for preds in where_*)
   4.604 +				   Calc("Poly.is'_polyexp",eval_is_polyexp"")],
   4.605 +		crls = e_rls, nrls = rls_p_34},
   4.606 +"Script SimplifyScript (t_::real) =                          " ^
   4.607 +"  ((Repeat((Try (Rewrite_Set klammern_aufloesen False)) @@  " ^
   4.608 +"           (Try (Rewrite_Set ordne_alphabetisch False)) @@  " ^
   4.609 +"           (Try (Rewrite_Set fasse_zusammen     False)) @@  " ^
   4.610 +"           (Try (Rewrite_Set verschoenere       False)))) t_)"
   4.611 +	       ));
   4.612 +
   4.613 +store_met
   4.614 +    (prep_met PolyMinus.thy "met_simp_poly_parenth_mult" [] e_metID
   4.615 +	      (["simplification","for_polynomials","with_parentheses_mult"],
   4.616 +	       [("#Given" ,["term t_"]),
   4.617 +		("#Where" ,["t_ is_polyexp"]),
   4.618 +		("#Find"  ,["normalform n_"])
   4.619 +		],
   4.620 +	       {rew_ord'="tless_true", rls' = e_rls, calc = [], srls = e_rls, 
   4.621 +		prls = append_rls "simplification_for_polynomials_prls" e_rls 
   4.622 +				  [(*for preds in where_*)
   4.623 +				   Calc("Poly.is'_polyexp",eval_is_polyexp"")],
   4.624 +		crls = e_rls, nrls = rls_p_34},
   4.625 +"Script SimplifyScript (t_::real) =                          " ^
   4.626 +"  ((Repeat((Try (Rewrite_Set klammern_ausmultiplizieren False)) @@ " ^
   4.627 +"           (Try (Rewrite_Set discard_parentheses        False)) @@ " ^
   4.628 +"           (Try (Rewrite_Set ordne_monome               False)) @@ " ^
   4.629 +"           (Try (Rewrite_Set klammern_aufloesen         False)) @@ " ^
   4.630 +"           (Try (Rewrite_Set ordne_alphabetisch         False)) @@ " ^
   4.631 +"           (Try (Rewrite_Set fasse_zusammen             False)) @@ " ^
   4.632 +"           (Try (Rewrite_Set verschoenere               False)))) t_)"
   4.633 +	       ));
   4.634 +
   4.635 +store_met
   4.636 +    (prep_met PolyMinus.thy "met_probe" [] e_metID
   4.637 +	      (["probe"],
   4.638 +	       [],
   4.639 +	       {rew_ord'="tless_true", rls' = e_rls, calc = [], srls = e_rls, 
   4.640 +		prls = Erls, crls = e_rls, nrls = Erls}, 
   4.641 +	       "empty_script"));
   4.642 +
   4.643 +store_met
   4.644 +    (prep_met PolyMinus.thy "met_probe_poly" [] e_metID
   4.645 +	      (["probe","fuer_polynom"],
   4.646 +	       [("#Given" ,["Pruefe e_", "mitWert ws_"]),
   4.647 +		("#Where" ,["e_ is_polyexp"]),
   4.648 +		("#Find"  ,["Geprueft p_"])
   4.649 +		],
   4.650 +	       {rew_ord'="tless_true", rls' = e_rls, calc = [], srls = e_rls, 
   4.651 +		prls = append_rls "prls_met_probe_bruch"
   4.652 +				  e_rls [(*for preds in where_*)
   4.653 +					 Calc ("Rational.is'_ratpolyexp", 
   4.654 +					       eval_is_ratpolyexp "")], 
   4.655 +		crls = e_rls, nrls = rechnen}, 
   4.656 +"Script ProbeScript (e_::bool) (ws_::bool list) = " ^
   4.657 +" (let e_ = Take e_;                              " ^
   4.658 +"      e_ = Substitute ws_ e_                     " ^
   4.659 +" in (Repeat((Try (Repeat (Calculate times))) @@  " ^
   4.660 +"            (Try (Repeat (Calculate plus ))) @@  " ^
   4.661 +"            (Try (Repeat (Calculate minus))))) e_)"
   4.662 +));
   4.663 +
   4.664 +store_met
   4.665 +    (prep_met PolyMinus.thy "met_probe_bruch" [] e_metID
   4.666 +	      (["probe","fuer_bruch"],
   4.667 +	       [("#Given" ,["Pruefe e_", "mitWert ws_"]),
   4.668 +		("#Where" ,["e_ is_ratpolyexp"]),
   4.669 +		("#Find"  ,["Geprueft p_"])
   4.670 +		],
   4.671 +	       {rew_ord'="tless_true", rls' = e_rls, calc = [], srls = e_rls, 
   4.672 +		prls = append_rls "prls_met_probe_bruch"
   4.673 +				  e_rls [(*for preds in where_*)
   4.674 +					 Calc ("Rational.is'_ratpolyexp", 
   4.675 +					       eval_is_ratpolyexp "")], 
   4.676 +		crls = e_rls, nrls = Erls}, 
   4.677 +	       "empty_script"));
   4.678 +*}
   4.679  
   4.680  end
   4.681  
     5.1 --- a/src/Tools/isac/Knowledge/Rational.thy	Thu Aug 26 10:03:53 2010 +0200
     5.2 +++ b/src/Tools/isac/Knowledge/Rational.thy	Thu Aug 26 18:15:30 2010 +0200
     5.3 @@ -5,44 +5,35 @@
     5.4  
     5.5     depends on Poly (and not on Atools), because 
     5.6     fractions with _normalized_ polynomials are canceled, added, etc.
     5.7 -
     5.8 -   use_thy_only"Knowledge/Rational";
     5.9 -   use_thy"../Knowledge/Rational";
    5.10 -   use_thy"Knowledge/Rational";
    5.11 -
    5.12 -   remove_thy"Rational";
    5.13 -   use_thy"Knowledge/Isac";
    5.14 -   use_thy_only"Knowledge/Rational";
    5.15 -
    5.16  *)
    5.17  
    5.18 -Rational = Poly +
    5.19 +theory Rational imports Poly begin
    5.20  
    5.21  consts
    5.22  
    5.23    is'_expanded   :: "real => bool" ("_ is'_expanded")     (*RL->Poly.thy*)
    5.24    is'_ratpolyexp :: "real => bool" ("_ is'_ratpolyexp") 
    5.25  
    5.26 -rules (*.not contained in Isabelle2002,
    5.27 -         stated as axioms, TODO: prove as theorems*)
    5.28 -
    5.29 -  mult_cross   "[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)"
    5.30 -  mult_cross1  "   b ~= 0            ==> (a / b = c    ) = (a     = b * c)"
    5.31 -  mult_cross2  "           d ~= 0    ==> (a     = c / d) = (a * d =     c)"
    5.32 -
    5.33 -  add_minus  "a + b - b = a"(*RL->Poly.thy*)
    5.34 -  add_minus1 "a - b + b = a"(*RL->Poly.thy*)
    5.35 -
    5.36 -  rat_mult                "a / b * (c / d) = a * c / (b * d)"(*?Isa02*) 
    5.37 -  rat_mult2               "a / b *  c      = a * c /  b     "(*?Isa02*)
    5.38 -
    5.39 -  rat_mult_poly_l         "c is_polyexp ==> c * (a / b) = c * a /  b"
    5.40 -  rat_mult_poly_r         "c is_polyexp ==> (a / b) * c = a * c /  b"
    5.41 +axioms (*.not contained in Isabelle2002,
    5.42 +          stated as axioms, TODO?: prove as theorems*)
    5.43 +
    5.44 +  mult_cross      "[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)"
    5.45 +  mult_cross1     "   b ~= 0            ==> (a / b = c    ) = (a     = b * c)"
    5.46 +  mult_cross2     "           d ~= 0    ==> (a     = c / d) = (a * d =     c)"
    5.47 +                  
    5.48 +  add_minus       "a + b - b = a"(*RL->Poly.thy*)
    5.49 +  add_minus1      "a - b + b = a"(*RL->Poly.thy*)
    5.50 +                  
    5.51 +  rat_mult        "a / b * (c / d) = a * c / (b * d)"(*?Isa02*) 
    5.52 +  rat_mult2       "a / b *  c      = a * c /  b     "(*?Isa02*)
    5.53 +
    5.54 +  rat_mult_poly_l "c is_polyexp ==> c * (a / b) = c * a /  b"
    5.55 +  rat_mult_poly_r "c is_polyexp ==> (a / b) * c = a * c /  b"
    5.56  
    5.57  (*real_times_divide1_eq .. Isa02*) 
    5.58    real_times_divide_1_eq  "-1    * (c / d) =-1 * c /      d "
    5.59 -  real_times_divide_num   "a is_const ==> \
    5.60 -	          	  \a     * (c / d) = a * c /      d "
    5.61 +  real_times_divide_num   "a is_const ==> 
    5.62 +	          	   a     * (c / d) = a * c /      d "
    5.63  
    5.64    real_mult_div_cancel2   "k ~= 0 ==> m * k / (n * k) = m / n"
    5.65  (*real_mult_div_cancel1   "k ~= 0 ==> k * m / (k * n) = m / n"..Isa02*)
    5.66 @@ -54,23 +45,3800 @@
    5.67    rat_power               "(a / b)^^^n = (a^^^n) / (b^^^n)"
    5.68  
    5.69  
    5.70 -  rat_add         "[| a is_const; b is_const; c is_const; d is_const |] ==> \
    5.71 -	          \a / c + b / d = (a * d + b * c) / (c * d)"
    5.72 -  rat_add_assoc   "[| a is_const; b is_const; c is_const; d is_const |] ==> \
    5.73 -	          \a / c +(b / d + e) = (a * d + b * c)/(d * c) + e"
    5.74 -  rat_add1        "[| a is_const; b is_const; c is_const |] ==> \
    5.75 -	          \a / c + b / c = (a + b) / c"
    5.76 -  rat_add1_assoc   "[| a is_const; b is_const; c is_const |] ==> \
    5.77 -	          \a / c + (b / c + e) = (a + b) / c + e"
    5.78 -  rat_add2        "[| a is_const; b is_const; c is_const |] ==> \
    5.79 -	          \a / c + b = (a + b * c) / c"
    5.80 -  rat_add2_assoc  "[| a is_const; b is_const; c is_const |] ==> \
    5.81 -	          \a / c + (b + e) = (a + b * c) / c + e"
    5.82 -  rat_add3        "[| a is_const; b is_const; c is_const |] ==> \
    5.83 -	          \a + b / c = (a * c + b) / c"
    5.84 -  rat_add3_assoc   "[| a is_const; b is_const; c is_const |] ==> \
    5.85 -	          \a + (b / c + e) = (a * c + b) / c + e"
    5.86 -
    5.87 -
    5.88 -
    5.89 +  rat_add         "[| a is_const; b is_const; c is_const; d is_const |] ==> 
    5.90 +	           a / c + b / d = (a * d + b * c) / (c * d)"
    5.91 +  rat_add_assoc   "[| a is_const; b is_const; c is_const; d is_const |] ==> 
    5.92 +	           a / c +(b / d + e) = (a * d + b * c)/(d * c) + e"
    5.93 +  rat_add1        "[| a is_const; b is_const; c is_const |] ==> 
    5.94 +	           a / c + b / c = (a + b) / c"
    5.95 +  rat_add1_assoc   "[| a is_const; b is_const; c is_const |] ==> 
    5.96 +	           a / c + (b / c + e) = (a + b) / c + e"
    5.97 +  rat_add2        "[| a is_const; b is_const; c is_const |] ==> 
    5.98 +	           a / c + b = (a + b * c) / c"
    5.99 +  rat_add2_assoc  "[| a is_const; b is_const; c is_const |] ==> 
   5.100 +	           a / c + (b + e) = (a + b * c) / c + e"
   5.101 +  rat_add3        "[| a is_const; b is_const; c is_const |] ==> 
   5.102 +	           a + b / c = (a * c + b) / c"
   5.103 +  rat_add3_assoc   "[| a is_const; b is_const; c is_const |] ==> 
   5.104 +	           a + (b / c + e) = (a * c + b) / c + e"
   5.105 +
   5.106 +text {*calculate in rationals: gcd, lcm, etc.
   5.107 +      (c) Stefan Karnel 2002
   5.108 +      Institute for Mathematics D and Institute for Software Technology, 
   5.109 +      TU-Graz SS 2002 *}
   5.110 +
   5.111 +text {* Remark on notions in the documentation below:
   5.112 +    referring to the remark on 'polynomials' in Poly.sml we use
   5.113 +    [2] 'polynomial' normalform (Polynom)
   5.114 +    [3] 'expanded_term' normalform (Ausmultiplizierter Term),
   5.115 +    where normalform [2] is a special case of [3], i.e. [3] implies [2].
   5.116 +    Instead of 
   5.117 +      'fraction with numerator and nominator both in normalform [2]'
   5.118 +      'fraction with numerator and nominator both in normalform [3]' 
   5.119 +    we say: 
   5.120 +      'fraction in normalform [2]'
   5.121 +      'fraction in normalform [3]' 
   5.122 +    or
   5.123 +      'fraction [2]'
   5.124 +      'fraction [3]'.
   5.125 +    a 'simple fraction' is a term with '/' as outmost operator and
   5.126 +    numerator and nominator in normalform [2] or [3].
   5.127 +*}
   5.128 +
   5.129 +ML {*
   5.130 +signature RATIONALI =
   5.131 +sig
   5.132 +  type mv_monom
   5.133 +  type mv_poly 
   5.134 +  val add_fraction_ : theory -> term -> (term * term list) option      
   5.135 +  val add_fraction_p_ : theory -> term -> (term * term list) option       
   5.136 +  val calculate_Rational : rls
   5.137 +  val calc_rat_erls:rls
   5.138 +  val cancel : rls
   5.139 +  val cancel_ : theory -> term -> (term * term list) option    
   5.140 +  val cancel_p : rls   
   5.141 +  val cancel_p_ : theory -> term -> (term * term list) option
   5.142 +  val common_nominator : rls              
   5.143 +  val common_nominator_ : theory -> term -> (term * term list) option
   5.144 +  val common_nominator_p : rls              
   5.145 +  val common_nominator_p_ : theory -> term -> (term * term list) option
   5.146 +  val eval_is_expanded : string -> 'a -> term -> theory -> 
   5.147 +			 (string * term) option                    
   5.148 +  val expanded2polynomial : term -> term option
   5.149 +  val factout_ : theory -> term -> (term * term list) option
   5.150 +  val factout_p_ : theory -> term -> (term * term list) option
   5.151 +  val is_expanded : term -> bool
   5.152 +  val is_polynomial : term -> bool
   5.153 +
   5.154 +  val mv_gcd : (int * int list) list -> mv_poly -> mv_poly
   5.155 +  val mv_lcm : mv_poly -> mv_poly -> mv_poly
   5.156 +
   5.157 +  val norm_expanded_rat_ : theory -> term -> (term * term list) option
   5.158 +(*WN0602.2.6.pull into struct !!!
   5.159 +  val norm_Rational : rls(*.normalizes an arbitrary rational term without
   5.160 +                           roots into a simple and canceled fraction
   5.161 +                           with normalform [2].*)
   5.162 +*)
   5.163 +(*val norm_rational_p : 19.10.02 missing FIXXXXXXXXXXXXME
   5.164 +      rls               (*.normalizes an rational term [2] without
   5.165 +                           roots into a simple and canceled fraction
   5.166 +                           with normalform [2].*)
   5.167 +*)
   5.168 +  val norm_rational_ : theory -> term -> (term * term list) option
   5.169 +  val polynomial2expanded : term -> term option
   5.170 +  val rational_erls : 
   5.171 +      rls             (*.evaluates an arbitrary rational term with numerals.*)
   5.172 +
   5.173 +(*WN0210???SK: fehlen Funktionen, die exportiert werden sollen ? *)
   5.174  end
   5.175 +
   5.176 +(*.**************************************************************************
   5.177 +survey on the functions
   5.178 +~~~~~~~~~~~~~~~~~~~~~~~
   5.179 + [2] 'polynomial'   :rls               | [3]'expanded_term':rls
   5.180 +--------------------:------------------+-------------------:-----------------
   5.181 + factout_p_         :                  | factout_          :
   5.182 + cancel_p_          :                  | cancel_           :
   5.183 +                    :cancel_p          |                   :cancel
   5.184 +--------------------:------------------+-------------------:-----------------
   5.185 + common_nominator_p_:                  | common_nominator_ :
   5.186 +                    :common_nominator_p|                   :common_nominator
   5.187 + add_fraction_p_    :                  | add_fraction_     :
   5.188 +--------------------:------------------+-------------------:-----------------
   5.189 +???SK                 :norm_rational_p   |                   :norm_rational
   5.190 +
   5.191 +This survey shows only the principal functions for reuse, and the identifiers 
   5.192 +of the rls exported. The list below shows some more useful functions.
   5.193 +
   5.194 +
   5.195 +conversion from Isabelle-term to internal representation
   5.196 +~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
   5.197 +
   5.198 +... BITTE FORTSETZEN ...
   5.199 +
   5.200 +polynomial2expanded = ...
   5.201 +expanded2polynomial = ...
   5.202 +
   5.203 +remark: polynomial2expanded o expanded2polynomial = I, 
   5.204 +        where 'o' is function chaining, and 'I' is identity WN0210???SK
   5.205 +
   5.206 +functions for greatest common divisor and canceling
   5.207 +~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
   5.208 +mv_gcd
   5.209 +factout_
   5.210 +factout_p_
   5.211 +cancel_
   5.212 +cancel_p_
   5.213 +
   5.214 +functions for least common multiple and addition of fractions
   5.215 +~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
   5.216 +mv_lcm
   5.217 +common_nominator_
   5.218 +common_nominator_p_
   5.219 +add_fraction_       (*.add 2 or more fractions.*)
   5.220 +add_fraction_p_     (*.add 2 or more fractions.*)
   5.221 +
   5.222 +functions for normalform of rationals
   5.223 +~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
   5.224 +WN0210???SK interne Funktionen f"ur norm_rational: 
   5.225 +          schaffen diese SML-Funktionen wirklich ganz allgemeine Terme ?
   5.226 +
   5.227 +norm_rational_
   5.228 +norm_expanded_rat_
   5.229 +
   5.230 +**************************************************************************.*)
   5.231 +
   5.232 +
   5.233 +(*##*)
   5.234 +structure RationalI : RATIONALI = 
   5.235 +struct 
   5.236 +(*##*)
   5.237 +
   5.238 +infix mem ins union; (*WN100819 updating to Isabelle2009-2*)
   5.239 +fun x mem [] = false
   5.240 +  | x mem (y :: ys) = x = y orelse x mem ys;
   5.241 +fun (x ins xs) = if x mem xs then xs else x :: xs;
   5.242 +fun xs union [] = xs
   5.243 +  | [] union ys = ys
   5.244 +  | (x :: xs) union ys = xs union (x ins ys);
   5.245 +
   5.246 +(*. gcd of integers .*)
   5.247 +(* die gcd Funktion von Isabelle funktioniert nicht richtig !!! *)
   5.248 +fun gcd_int a b = if b=0 then a
   5.249 +		  else gcd_int b (a mod b);
   5.250 +
   5.251 +(*. univariate polynomials (uv) .*)
   5.252 +(*. univariate polynomials are represented as a list 
   5.253 +    of the coefficent in reverse maximum degree order .*)
   5.254 +(*. 5 * x^5 + 4 * x^3 + 2 * x^2 + x + 19 => [19,1,2,4,0,5] .*)
   5.255 +type uv_poly = int list;
   5.256 +
   5.257 +(*. adds two uv polynomials .*)
   5.258 +fun uv_mod_add_poly ([]:uv_poly,p2:uv_poly) = p2:uv_poly 
   5.259 +  | uv_mod_add_poly (p1,[]) = p1
   5.260 +  | uv_mod_add_poly (x::p1,y::p2) = (x+y)::(uv_mod_add_poly(p1,p2)); 
   5.261 +
   5.262 +(*. multiplies a uv polynomial with a skalar s .*)
   5.263 +fun uv_mod_smul_poly ([]:uv_poly,s:int) = []:uv_poly 
   5.264 +  | uv_mod_smul_poly (x::p,s) = (x*s)::(uv_mod_smul_poly(p,s)); 
   5.265 +
   5.266 +(*. calculates the remainder of a polynomial divided by a skalar s .*)
   5.267 +fun uv_mod_rem_poly ([]:uv_poly,s) = []:uv_poly 
   5.268 +  | uv_mod_rem_poly (x::p,s) = (x mod s)::(uv_mod_smul_poly(p,s)); 
   5.269 +
   5.270 +(*. calculates the degree of a uv polynomial .*)
   5.271 +fun uv_mod_deg ([]:uv_poly) = 0  
   5.272 +  | uv_mod_deg p = length(p)-1;
   5.273 +
   5.274 +(*. calculates the remainder of x/p and represents it as 
   5.275 +    value between -p/2 and p/2 .*)
   5.276 +fun uv_mod_mod2(x,p)=
   5.277 +    let
   5.278 +	val y=(x mod p);
   5.279 +    in
   5.280 +	if (y)>(p div 2) then (y)-p else 
   5.281 +	    (
   5.282 +	     if (y)<(~p div 2) then p+(y) else (y)
   5.283 +	     )
   5.284 +    end;
   5.285 +
   5.286 +(*.calculates the remainder for each element of a integer list divided by p.*)  
   5.287 +fun uv_mod_list_modp [] p = [] 
   5.288 +  | uv_mod_list_modp (x::xs) p = (uv_mod_mod2(x,p))::(uv_mod_list_modp xs p);
   5.289 +
   5.290 +(*. appends an integer at the end of a integer list .*)
   5.291 +fun uv_mod_null (p1:int list,0) = p1 
   5.292 +  | uv_mod_null (p1:int list,n1:int) = uv_mod_null(p1,n1-1) @ [0];
   5.293 +
   5.294 +(*. uv polynomial division, result is (quotient, remainder) .*)
   5.295 +(*. only for uv_mod_divides .*)
   5.296 +(* FIXME: Division von x^9+x^5+1 durch x-1000 funktioniert nicht,
   5.297 +   integer zu klein  *)
   5.298 +fun uv_mod_pdiv (p1:uv_poly) ([]:uv_poly) = 
   5.299 +    raise error ("RATIONALS_UV_MOD_PDIV_EXCEPTION: division by zero")
   5.300 +  | uv_mod_pdiv p1 [x] = 
   5.301 +    let
   5.302 +	val xs=ref [];
   5.303 +    in
   5.304 +	if x<>0 then 
   5.305 +	    (
   5.306 +	     xs:=(uv_mod_rem_poly(p1,x));
   5.307 +	     while length(!xs)>0 andalso hd(!xs)=0 do xs:=tl(!xs)
   5.308 +	     )
   5.309 +	else raise error ("RATIONALS_UV_MOD_PDIV_EXCEPTION: division by zero");
   5.310 +	([]:uv_poly,!xs:uv_poly)
   5.311 +    end
   5.312 +  | uv_mod_pdiv p1 p2 =  
   5.313 +    let
   5.314 +	val n= uv_mod_deg(p2);
   5.315 +	val m= ref (uv_mod_deg(p1));
   5.316 +	val p1'=ref (rev(p1));
   5.317 +	val p2'=(rev(p2));
   5.318 +	val lc2=hd(p2');
   5.319 +	val q=ref [];
   5.320 +	val c=ref 0;
   5.321 +	val output=ref ([],[]);
   5.322 +    in
   5.323 +	(
   5.324 +	 if (!m)=0 orelse p2=[0] 
   5.325 +         then raise error ("RATIONALS_UV_MOD_PDIV_EXCEPTION: Division by zero") 
   5.326 +	 else
   5.327 +	     (
   5.328 +	      if (!m)<n then 
   5.329 +		  (
   5.330 +		   output:=([0],p1) 
   5.331 +		   ) 
   5.332 +	      else
   5.333 +		  (
   5.334 +		   while (!m)>=n do
   5.335 +		       (
   5.336 +			c:=hd(!p1') div hd(p2');
   5.337 +			if !c<>0 then
   5.338 +			    (
   5.339 +			     p1':=uv_mod_add_poly(!p1',uv_mod_null(uv_mod_smul_poly(p2',~(!c)),!m-n));
   5.340 +			     while length(!p1')>0 andalso hd(!p1')=0  do p1':= tl(!p1');
   5.341 +			     m:=uv_mod_deg(!p1')
   5.342 +			     )
   5.343 +			else m:=0
   5.344 +			);
   5.345 +    		   output:=(rev(!q),rev(!p1'))
   5.346 +		   )
   5.347 +	      );
   5.348 +	     !output
   5.349 +	 )
   5.350 +    end;
   5.351 +
   5.352 +(*. divides p1 by p2 in Zp .*)
   5.353 +fun uv_mod_pdivp (p1:uv_poly) (p2:uv_poly) p =  
   5.354 +    let
   5.355 +	val n=uv_mod_deg(p2);
   5.356 +	val m=ref (uv_mod_deg(uv_mod_list_modp p1 p));
   5.357 +	val p1'=ref (rev(p1));
   5.358 +	val p2'=(rev(uv_mod_list_modp p2 p));
   5.359 +	val lc2=hd(p2');
   5.360 +	val q=ref [];
   5.361 +	val c=ref 0;
   5.362 +	val output=ref ([],[]);
   5.363 +    in
   5.364 +	(
   5.365 +	 if (!m)=0 orelse p2=[0] then raise error ("RATIONALS_UV_MOD_PDIVP_EXCEPTION: Division by zero") 
   5.366 +	 else
   5.367 +	     (
   5.368 +	      if (!m)<n then 
   5.369 +		  (
   5.370 +		   output:=([0],p1) 
   5.371 +		   ) 
   5.372 +	      else
   5.373 +		  (
   5.374 +		   while (!m)>=n do
   5.375 +		       (
   5.376 +			c:=uv_mod_mod2(hd(!p1')*(power lc2 1), p);
   5.377 +			q:=(!c)::(!q);
   5.378 +			p1':=uv_mod_list_modp(tl(uv_mod_add_poly(uv_mod_smul_poly(!p1',lc2),
   5.379 +								  uv_mod_smul_poly(uv_mod_smul_poly(p2',hd(!p1')),~1)))) p;
   5.380 +			m:=(!m)-1
   5.381 +			);
   5.382 +		   
   5.383 +		   while !p1'<>[] andalso hd(!p1')=0 do
   5.384 +		       (
   5.385 +			p1':=tl(!p1')
   5.386 +			);
   5.387 +
   5.388 +    		   output:=(rev(uv_mod_list_modp (!q) (p)),rev(!p1'))
   5.389 +		   )
   5.390 +	      );
   5.391 +	     !output:uv_poly * uv_poly
   5.392 +	 )
   5.393 +    end;
   5.394 +
   5.395 +(*. calculates the remainder of p1/p2 .*)
   5.396 +fun uv_mod_prest (p1:uv_poly) ([]:uv_poly) = raise error("UV_MOD_PREST_EXCEPTION: Division by zero") 
   5.397 +  | uv_mod_prest [] p2 = []:uv_poly
   5.398 +  | uv_mod_prest p1 p2 = (#2(uv_mod_pdiv p1 p2));
   5.399 +
   5.400 +(*. calculates the remainder of p1/p2 in Zp .*)
   5.401 +fun uv_mod_prestp (p1:uv_poly) ([]:uv_poly) p= raise error("UV_MOD_PRESTP_EXCEPTION: Division by zero") 
   5.402 +  | uv_mod_prestp [] p2 p= []:uv_poly 
   5.403 +  | uv_mod_prestp p1 p2 p = #2(uv_mod_pdivp p1 p2 p); 
   5.404 +
   5.405 +(*. calculates the content of a uv polynomial .*)
   5.406 +fun uv_mod_cont ([]:uv_poly) = 0  
   5.407 +  | uv_mod_cont (x::p)= gcd_int x (uv_mod_cont(p));
   5.408 +
   5.409 +(*. divides each coefficient of a uv polynomial by y .*)
   5.410 +fun uv_mod_div_list (p:uv_poly,0) = raise error("UV_MOD_DIV_LIST_EXCEPTION: Division by zero") 
   5.411 +  | uv_mod_div_list ([],y)   = []:uv_poly
   5.412 +  | uv_mod_div_list (x::p,y) = (x div y)::uv_mod_div_list(p,y); 
   5.413 +
   5.414 +(*. calculates the primitiv part of a uv polynomial .*)
   5.415 +fun uv_mod_pp ([]:uv_poly) = []:uv_poly
   5.416 +  | uv_mod_pp p =  
   5.417 +    let
   5.418 +	val c=ref 0;
   5.419 +    in
   5.420 +	(
   5.421 +	 c:=uv_mod_cont(p);
   5.422 +	 
   5.423 +	 if !c=0 then raise error ("RATIONALS_UV_MOD_PP_EXCEPTION: content is 0")
   5.424 +	 else uv_mod_div_list(p,!c)
   5.425 +	)
   5.426 +    end;
   5.427 +
   5.428 +(*. gets the leading coefficient of a uv polynomial .*)
   5.429 +fun uv_mod_lc ([]:uv_poly) = 0 
   5.430 +  | uv_mod_lc p  = hd(rev(p)); 
   5.431 +
   5.432 +(*. calculates the euklidean polynomial remainder sequence in Zp .*)
   5.433 +fun uv_mod_prs_euklid_p(p1:uv_poly,p2:uv_poly,p)= 
   5.434 +    let
   5.435 +	val f =ref [];
   5.436 +	val f'=ref p2;
   5.437 +	val fi=ref [];
   5.438 +    in
   5.439 +	( 
   5.440 +	 f:=p2::p1::[]; 
   5.441 + 	 while uv_mod_deg(!f')>0 do
   5.442 +	     (
   5.443 +	      f':=uv_mod_prestp (hd(tl(!f))) (hd(!f)) p;
   5.444 +	      if (!f')<>[] then 
   5.445 +		  (
   5.446 +		   fi:=(!f');
   5.447 +		   f:=(!fi)::(!f)
   5.448 +		   )
   5.449 +	      else ()
   5.450 +	      );
   5.451 +	      (!f)
   5.452 +	 
   5.453 +	 )
   5.454 +    end;
   5.455 +
   5.456 +(*. calculates the gcd of p1 and p2 in Zp .*)
   5.457 +fun uv_mod_gcd_modp ([]:uv_poly) (p2:uv_poly) p = p2:uv_poly 
   5.458 +  | uv_mod_gcd_modp p1 [] p= p1
   5.459 +  | uv_mod_gcd_modp p1 p2 p=
   5.460 +    let
   5.461 +	val p1'=ref[];
   5.462 +	val p2'=ref[];
   5.463 +	val pc=ref[];
   5.464 +	val g=ref [];
   5.465 +	val d=ref 0;
   5.466 +	val prs=ref [];
   5.467 +    in
   5.468 +	(
   5.469 +	 if uv_mod_deg(p1)>=uv_mod_deg(p2) then
   5.470 +	     (
   5.471 +	      p1':=uv_mod_list_modp (uv_mod_pp(p1)) p;
   5.472 +	      p2':=uv_mod_list_modp (uv_mod_pp(p2)) p
   5.473 +	      )
   5.474 +	 else 
   5.475 +	     (
   5.476 +	      p1':=uv_mod_list_modp (uv_mod_pp(p2)) p;
   5.477 +	      p2':=uv_mod_list_modp (uv_mod_pp(p1)) p
   5.478 +	      );
   5.479 +	 d:=uv_mod_mod2((gcd_int (uv_mod_cont(p1))) (uv_mod_cont(p2)), p) ;
   5.480 +	 if !d>(p div 2) then d:=(!d)-p else ();
   5.481 +	 
   5.482 +	 prs:=uv_mod_prs_euklid_p(!p1',!p2',p);
   5.483 +
   5.484 +	 if hd(!prs)=[] then pc:=hd(tl(!prs))
   5.485 +	 else pc:=hd(!prs);
   5.486 +
   5.487 +	 g:=uv_mod_smul_poly(uv_mod_pp(!pc),!d);
   5.488 +	 !g
   5.489 +	 )
   5.490 +    end;
   5.491 +
   5.492 +(*. calculates the minimum of two real values x and y .*)
   5.493 +fun uv_mod_r_min(x,y):BasisLibrary.Real.real = if x>y then y else x;
   5.494 +
   5.495 +(*. calculates the minimum of two integer values x and y .*)
   5.496 +fun uv_mod_min(x,y) = if x>y then y else x;
   5.497 +
   5.498 +(*. adds the squared values of a integer list .*)
   5.499 +fun uv_mod_add_qu [] = 0.0 
   5.500 +  | uv_mod_add_qu (x::p) =  BasisLibrary.Real.fromInt(x)*BasisLibrary.Real.fromInt(x) + uv_mod_add_qu p;
   5.501 +
   5.502 +(*. calculates the euklidean norm .*)
   5.503 +fun uv_mod_norm ([]:uv_poly) = 0.0
   5.504 +  | uv_mod_norm p = Math.sqrt(uv_mod_add_qu(p));
   5.505 +
   5.506 +(*. multipies two values a and b .*)
   5.507 +fun uv_mod_multi a b = a * b;
   5.508 +
   5.509 +(*. decides if x is a prim, the list contains all primes which are lower then x .*)
   5.510 +fun uv_mod_prim(x,[])= false 
   5.511 +  | uv_mod_prim(x,[y])=if ((x mod y) <> 0) then true
   5.512 +		else false
   5.513 +  | uv_mod_prim(x,y::ys) = if uv_mod_prim(x,[y])
   5.514 +			then 
   5.515 +			    if uv_mod_prim(x,ys) then true 
   5.516 +			    else false
   5.517 +		    else false;
   5.518 +
   5.519 +(*. gets the first prime, which is greater than p and does not divide g .*)
   5.520 +fun uv_mod_nextprime(g,p)= 
   5.521 +    let
   5.522 +	val list=ref [2];
   5.523 +	val exit=ref 0;
   5.524 +	val i = ref 2
   5.525 +    in
   5.526 +	while (!i<p) do (* calculates the primes lower then p *)
   5.527 +	    (
   5.528 +	     if uv_mod_prim(!i,!list) then
   5.529 +		 (
   5.530 +		  if (p mod !i <> 0)
   5.531 +		      then
   5.532 +			  (
   5.533 +			   list:= (!i)::(!list);
   5.534 +			   i:= (!i)+1
   5.535 +			   )
   5.536 +		  else i:=(!i)+1
   5.537 +		  )
   5.538 +	     else i:= (!i)+1
   5.539 +		 );
   5.540 +	    i:=(p+1);
   5.541 +	    while (!exit=0) do   (* calculate next prime which does not divide g *)
   5.542 +	    (
   5.543 +	     if uv_mod_prim(!i,!list) then
   5.544 +		 (
   5.545 +		  if (g mod !i <> 0)
   5.546 +		      then
   5.547 +			  (
   5.548 +			   list:= (!i)::(!list);
   5.549 +			   exit:= (!i)
   5.550 +			   )
   5.551 +		  else i:=(!i)+1
   5.552 +		      )
   5.553 +	     else i:= (!i)+1
   5.554 +		 ); 
   5.555 +	    !exit
   5.556 +    end;
   5.557 +
   5.558 +(*. decides if p1 is a factor of p2 in Zp .*)
   5.559 +fun uv_mod_dividesp ([]:uv_poly) (p2:uv_poly) p= raise error("UV_MOD_DIVIDESP: Division by zero") 
   5.560 +  | uv_mod_dividesp p1 p2 p= if uv_mod_prestp p2 p1 p = [] then true else false;
   5.561 +
   5.562 +(*. decides if p1 is a factor of p2 .*)
   5.563 +fun uv_mod_divides ([]:uv_poly) (p2:uv_poly) = raise error("UV_MOD_DIVIDES: Division by zero")
   5.564 +  | uv_mod_divides p1 p2 = if uv_mod_prest p2 p1  = [] then true else false;
   5.565 +
   5.566 +(*. chinese remainder algorithm .*)
   5.567 +fun uv_mod_cra2(r1,r2,m1,m2)=     
   5.568 +    let 
   5.569 +	val c=ref 0;
   5.570 +	val r1'=ref 0;
   5.571 +	val d=ref 0;
   5.572 +	val a=ref 0;
   5.573 +    in
   5.574 +	(
   5.575 +	 while (uv_mod_mod2((!c)*m1,m2))<>1 do 
   5.576 +	     (
   5.577 +	      c:=(!c)+1
   5.578 +	      );
   5.579 +	 r1':= uv_mod_mod2(r1,m1);
   5.580 +	 d:=uv_mod_mod2(((r2-(!r1'))*(!c)),m2);
   5.581 +	 !r1'+(!d)*m1    
   5.582 +	 )
   5.583 +    end;
   5.584 +
   5.585 +(*. applies the chinese remainder algorithmen to the coefficients of x1 and x2 .*)
   5.586 +fun uv_mod_cra_2 ([],[],m1,m2) = [] 
   5.587 +  | uv_mod_cra_2 ([],x2,m1,m2) = raise error("UV_MOD_CRA_2_EXCEPTION: invalid call x1")
   5.588 +  | uv_mod_cra_2 (x1,[],m1,m2) = raise error("UV_MOD_CRA_2_EXCEPTION: invalid call x2")
   5.589 +  | 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));
   5.590 +
   5.591 +(*. calculates the gcd of two uv polynomials p1' and p2' with the modular algorithm .*)
   5.592 +fun uv_mod_gcd (p1':uv_poly) (p2':uv_poly) =
   5.593 +    let 
   5.594 +	val p1=ref (uv_mod_pp(p1'));
   5.595 +	val p2=ref (uv_mod_pp(p2'));
   5.596 +	val c=gcd_int (uv_mod_cont(p1')) (uv_mod_cont(p2'));
   5.597 +	val temp=ref [];
   5.598 +	val cp=ref [];
   5.599 +	val qp=ref [];
   5.600 +	val q=ref[];
   5.601 +	val pn=ref 0;
   5.602 +	val d=ref 0;
   5.603 +	val g1=ref 0;
   5.604 +	val p=ref 0;    
   5.605 +	val m=ref 0;
   5.606 +	val exit=ref 0;
   5.607 +	val i=ref 1;
   5.608 +    in
   5.609 +	if length(!p1)>length(!p2) then ()
   5.610 +	else 
   5.611 +	    (
   5.612 +	     temp:= !p1;
   5.613 +	     p1:= !p2;
   5.614 +	     p2:= !temp
   5.615 +	     );
   5.616 +
   5.617 +	 
   5.618 +	d:=gcd_int (uv_mod_lc(!p1)) (uv_mod_lc(!p2));
   5.619 +	g1:=uv_mod_lc(!p1)*uv_mod_lc(!p2);
   5.620 +	p:=4;
   5.621 +	
   5.622 +	m:=BasisLibrary.Real.ceil(2.0 *   
   5.623 +				  BasisLibrary.Real.fromInt(!d) *
   5.624 +				  BasisLibrary.Real.fromInt(power 2 (uv_mod_min(uv_mod_deg(!p2),uv_mod_deg(!p1)))) *  
   5.625 +				  BasisLibrary.Real.fromInt(!d) * 
   5.626 +				  uv_mod_r_min(uv_mod_norm(!p1) / BasisLibrary.Real.fromInt(abs(uv_mod_lc(!p1))),
   5.627 +					uv_mod_norm(!p2) / BasisLibrary.Real.fromInt(abs(uv_mod_lc(!p2))))); 
   5.628 +
   5.629 +	while (!exit=0) do  
   5.630 +	    (
   5.631 +	     p:=uv_mod_nextprime(!d,!p);
   5.632 +	     cp:=(uv_mod_gcd_modp (uv_mod_list_modp(!p1) (!p)) (uv_mod_list_modp(!p2) (!p)) (!p)) ;
   5.633 +	     if abs(uv_mod_lc(!cp))<>1 then  (* leading coefficient = 1 ? *)
   5.634 +		 (
   5.635 +		  i:=1;
   5.636 +		  while (!i)<(!p) andalso (abs(uv_mod_mod2((uv_mod_lc(!cp)*(!i)),(!p)))<>1) do
   5.637 +		      (
   5.638 +		       i:=(!i)+1
   5.639 +		       );
   5.640 +		      cp:=uv_mod_list_modp (map (uv_mod_multi (!i)) (!cp)) (!p) 
   5.641 +		  )
   5.642 +	     else ();
   5.643 +
   5.644 +	     qp:= ((map (uv_mod_multi (uv_mod_mod2(!d,!p)))) (!cp));
   5.645 +
   5.646 +	     if uv_mod_deg(!qp)=0 then (q:=[1]; exit:=1) else ();
   5.647 +
   5.648 +	     pn:=(!p);
   5.649 +	     q:=(!qp);
   5.650 +
   5.651 +	     while !pn<= !m andalso !m>(!p) andalso !exit=0 do
   5.652 +		 (
   5.653 +		  p:=uv_mod_nextprime(!d,!p);
   5.654 + 		  cp:=(uv_mod_gcd_modp (uv_mod_list_modp(!p1) (!p)) (uv_mod_list_modp(!p2) (!p)) (!p)); 
   5.655 + 		  if uv_mod_lc(!cp)<>1 then  (* leading coefficient = 1 ? *)
   5.656 + 		      (
   5.657 + 		       i:=1;
   5.658 + 		       while (!i)<(!p) andalso ((uv_mod_mod2((uv_mod_lc(!q)*(!i)),(!p)))<>1) do
   5.659 + 			   (
   5.660 + 			    i:=(!i)+1
   5.661 +		           );
   5.662 +		       cp:=uv_mod_list_modp (map (uv_mod_multi (!i)) (!cp)) (!p)
   5.663 + 		      )
   5.664 + 		  else ();    
   5.665 + 		 
   5.666 +		  qp:=uv_mod_list_modp ((map (uv_mod_multi (uv_mod_mod2(!d,!p)))) (!cp)  ) (!p);
   5.667 + 		  if uv_mod_deg(!qp)>uv_mod_deg(!q) then
   5.668 + 		      (
   5.669 + 		       (*print("degree to high!!!\n")*)
   5.670 + 		       )
   5.671 + 		  else
   5.672 + 		      (
   5.673 + 		       if uv_mod_deg(!qp)=uv_mod_deg(!q) then
   5.674 + 			   (
   5.675 + 			    q:=uv_mod_cra_2(!q,!qp,!pn,!p);
   5.676 +			    pn:=(!pn) * !p;
   5.677 +			    q:=uv_mod_pp(uv_mod_list_modp (!q) (!pn)); (* found already gcd ? *)
   5.678 +			    if (uv_mod_divides (!q) (p1')) andalso (uv_mod_divides (!q) (p2')) then (exit:=1) else ()
   5.679 +		 	    )
   5.680 +		       else
   5.681 +			   (
   5.682 +			    if  uv_mod_deg(!qp)<uv_mod_deg(!q) then
   5.683 +				(
   5.684 +				 pn:= !p;
   5.685 +				 q:= !qp
   5.686 +				 )
   5.687 +			    else ()
   5.688 +			    )
   5.689 +		       )
   5.690 +		  );
   5.691 + 	     q:=uv_mod_pp(uv_mod_list_modp (!q) (!pn));
   5.692 +	     if (uv_mod_divides (!q) (p1')) andalso (uv_mod_divides (!q) (p2')) then exit:=1 else ()
   5.693 +	     );
   5.694 +	    uv_mod_smul_poly(!q,c):uv_poly
   5.695 +    end;
   5.696 +
   5.697 +(*. multivariate polynomials .*)
   5.698 +(*. multivariate polynomials are represented as a list of the pairs, 
   5.699 + first is the coefficent and the second is a list of the exponents .*)
   5.700 +(*. 5 * x^5 * y^3 + 4 * x^3 * z^2 + 2 * x^2 * y * z^3 - z - 19 
   5.701 + => [(5,[5,3,0]),(4,[3,0,2]),(2,[2,1,3]),(~1,[0,0,1]),(~19,[0,0,0])] .*)
   5.702 +
   5.703 +(*. global variables .*)
   5.704 +(*. order indicators .*)
   5.705 +val LEX_=0; (* lexicographical term order *)
   5.706 +val GGO_=1; (* greatest degree order *)
   5.707 +
   5.708 +(*. datatypes for internal representation.*)
   5.709 +type mv_monom = (int *      (*.coefficient or the monom.*)
   5.710 +		 int list); (*.list of exponents)      .*)
   5.711 +fun mv_monom2str (i, is) = "("^ int2str i^"," ^ ints2str' is ^ ")";
   5.712 +
   5.713 +type mv_poly = mv_monom list; 
   5.714 +fun mv_poly2str p = (strs2str' o (map mv_monom2str)) p;
   5.715 +
   5.716 +(*. help function for monom_greater and geq .*)
   5.717 +fun mv_mg_hlp([]) = EQUAL 
   5.718 +  | mv_mg_hlp(x::list)=if x<0 then LESS
   5.719 +		    else if x>0 then GREATER
   5.720 +			 else mv_mg_hlp(list);
   5.721 +
   5.722 +(*. adds a list of values .*)
   5.723 +fun mv_addlist([]) = 0
   5.724 +  | mv_addlist(p1) = hd(p1)+mv_addlist(tl(p1));
   5.725 +			   
   5.726 +(*. tests if the monomial M1 is greater as the monomial M2 and returns a boolean value .*)
   5.727 +(*. 2 orders are implemented LEX_/GGO_ (lexigraphical/greatest degree order) .*)
   5.728 +fun mv_monom_greater((M1x,M1l):mv_monom,(M2x,M2l):mv_monom,order)=
   5.729 +    if order=LEX_ then
   5.730 +	( 
   5.731 +	 if length(M1l)<>length(M2l) then raise error ("RATIONALS_MV_MONOM_GREATER_EXCEPTION: Order error")
   5.732 +	 else if (mv_mg_hlp((map op- (M1l~~M2l)))<>GREATER) then false else true
   5.733 +	     )
   5.734 +    else
   5.735 +	if order=GGO_ then
   5.736 +	    ( 
   5.737 +	     if length(M1l)<>length(M2l) then raise error ("RATIONALS_MV_MONOM_GREATER_EXCEPTION: Order error")
   5.738 +	     else 
   5.739 +		 if mv_addlist(M1l)=mv_addlist(M2l)  then if (mv_mg_hlp((map op- (M1l~~M2l)))<>GREATER) then false else true
   5.740 +		 else if mv_addlist(M1l)>mv_addlist(M2l) then true else false
   5.741 +	     )
   5.742 +	else raise error ("RATIONALS_MV_MONOM_GREATER_EXCEPTION: Wrong Order");
   5.743 +		   
   5.744 +(*. tests if the monomial X is greater as the monomial Y and returns a order value (GREATER,EQUAL,LESS) .*)
   5.745 +(*. 2 orders are implemented LEX_/GGO_ (lexigraphical/greatest degree order) .*)
   5.746 +fun mv_geq order ((x1,x):mv_monom,(x2,y):mv_monom) =
   5.747 +let 
   5.748 +    val temp=ref EQUAL;
   5.749 +in
   5.750 +    if order=LEX_ then
   5.751 +	(
   5.752 +	 if length(x)<>length(y) then 
   5.753 +	     raise error ("RATIONALS_MV_GEQ_EXCEPTION: Order error")
   5.754 +	 else 
   5.755 +	     (
   5.756 +	      temp:=mv_mg_hlp((map op- (x~~y)));
   5.757 +	      if !temp=EQUAL then 
   5.758 +		  ( if x1=x2 then EQUAL 
   5.759 +		    else if x1>x2 then GREATER
   5.760 +			 else LESS
   5.761 +			     )
   5.762 +	      else (!temp)
   5.763 +	      )
   5.764 +	     )
   5.765 +    else 
   5.766 +	if order=GGO_ then 
   5.767 +	    (
   5.768 +	     if length(x)<>length(y) then 
   5.769 +	      raise error ("RATIONALS_MV_GEQ_EXCEPTION: Order error")
   5.770 +	     else 
   5.771 +		 if mv_addlist(x)=mv_addlist(y) then 
   5.772 +		     (mv_mg_hlp((map op- (x~~y))))
   5.773 +		 else if mv_addlist(x)>mv_addlist(y) then GREATER else LESS
   5.774 +		     )
   5.775 +	else raise error ("RATIONALS_MV_GEQ_EXCEPTION: Wrong Order")
   5.776 +end;
   5.777 +
   5.778 +(*. cuts the first variable from a polynomial .*)
   5.779 +fun mv_cut([]:mv_poly)=[]:mv_poly
   5.780 +  | mv_cut((x,[])::list) = raise error ("RATIONALS_MV_CUT_EXCEPTION: Invalid list ")
   5.781 +  | mv_cut((x,y::ys)::list)=(x,ys)::mv_cut(list);
   5.782 +	    
   5.783 +(*. leading power product .*)
   5.784 +fun mv_lpp([]:mv_poly,order)  = []
   5.785 +  | mv_lpp([(x,y)],order) = y
   5.786 +  | mv_lpp(p1,order)  = #2(hd(rev(sort (mv_geq order) p1)));
   5.787 +    
   5.788 +(*. leading monomial .*)
   5.789 +fun mv_lm([]:mv_poly,order)  = (0,[]):mv_monom
   5.790 +  | mv_lm([x],order) = x 
   5.791 +  | mv_lm(p1,order)  = hd(rev(sort (mv_geq order) p1));
   5.792 +    
   5.793 +(*. leading coefficient in term order .*)
   5.794 +fun mv_lc2([]:mv_poly,order)  = 0
   5.795 +  | mv_lc2([(x,y)],order) = x
   5.796 +  | mv_lc2(p1,order)  = #1(hd(rev(sort (mv_geq order) p1)));
   5.797 +
   5.798 +
   5.799 +(*. reverse the coefficients in mv polynomial .*)
   5.800 +fun mv_rev_to([]:mv_poly) = []:mv_poly
   5.801 +  | mv_rev_to((c,e)::xs) = (c,rev(e))::mv_rev_to(xs);
   5.802 +
   5.803 +(*. leading coefficient in reverse term order .*)
   5.804 +fun mv_lc([]:mv_poly,order)  = []:mv_poly 
   5.805 +  | mv_lc([(x,y)],order) = mv_rev_to(mv_cut(mv_rev_to([(x,y)])))
   5.806 +  | mv_lc(p1,order)  = 
   5.807 +    let
   5.808 +	val p1o=ref (rev(sort (mv_geq order) (mv_rev_to(p1))));
   5.809 +	val lp=hd(#2(hd(!p1o)));
   5.810 +	val lc=ref [];
   5.811 +    in
   5.812 +	(
   5.813 +	 while (length(!p1o)>0 andalso hd(#2(hd(!p1o)))=lp) do
   5.814 +	     (
   5.815 +	      lc:=hd(mv_cut([hd(!p1o)]))::(!lc);
   5.816 +	      p1o:=tl(!p1o)
   5.817 +	      );
   5.818 +	 if !lc=[] then raise error ("RATIONALS_MV_LC_EXCEPTION: lc is empty") else ();
   5.819 +	 mv_rev_to(!lc)
   5.820 +	 )
   5.821 +    end;
   5.822 +
   5.823 +(*. compares two powerproducts .*)
   5.824 +fun mv_monom_equal((_,xlist):mv_monom,(_,ylist):mv_monom) = (foldr and_) (((map op=) (xlist~~ylist)),true);
   5.825 +    
   5.826 +(*. help function for mv_add .*)
   5.827 +fun mv_madd([]:mv_poly,[]:mv_poly,order) = []:mv_poly
   5.828 +  | mv_madd([(0,_)],p2,order) = p2
   5.829 +  | mv_madd(p1,[(0,_)],order) = p1  
   5.830 +  | mv_madd([],p2,order) = p2
   5.831 +  | mv_madd(p1,[],order) = p1
   5.832 +  | mv_madd(p1,p2,order) = 
   5.833 +    (
   5.834 +     if mv_monom_greater(hd(p1),hd(p2),order) 
   5.835 +	 then hd(p1)::mv_madd(tl(p1),p2,order)
   5.836 +     else if mv_monom_equal(hd(p1),hd(p2)) 
   5.837 +	      then if mv_lc2(p1,order)+mv_lc2(p2,order)<>0 
   5.838 +		       then (mv_lc2(p1,order)+mv_lc2(p2,order),mv_lpp(p1,order))::mv_madd(tl(p1),tl(p2),order)
   5.839 +		   else mv_madd(tl(p1),tl(p2),order)
   5.840 +	  else hd(p2)::mv_madd(p1,tl(p2),order)
   5.841 +	      )
   5.842 +	      
   5.843 +(*. adds two multivariate polynomials .*)
   5.844 +fun mv_add([]:mv_poly,p2:mv_poly,order) = p2
   5.845 +  | mv_add(p1,[],order) = p1
   5.846 +  | mv_add(p1,p2,order) = mv_madd(rev(sort (mv_geq order) p1),rev(sort (mv_geq order) p2), order);
   5.847 +
   5.848 +(*. monom multiplication .*)
   5.849 +fun mv_mmul((x1,y1):mv_monom,(x2,y2):mv_monom)=(x1*x2,(map op+) (y1~~y2)):mv_monom;
   5.850 +
   5.851 +(*. deletes all monomials with coefficient 0 .*)
   5.852 +fun mv_shorten([]:mv_poly,order) = []:mv_poly
   5.853 +  | mv_shorten(x::xs,order)=mv_madd([x],mv_shorten(xs,order),order);
   5.854 +
   5.855 +(*. zeros a list .*)
   5.856 +fun mv_null2([])=[]
   5.857 +  | mv_null2(x::l)=0::mv_null2(l);
   5.858 +
   5.859 +(*. multiplies two multivariate polynomials .*)
   5.860 +fun mv_mul([]:mv_poly,[]:mv_poly,_) = []:mv_poly
   5.861 +  | mv_mul([],y::p2,_) = [(0,mv_null2(#2(y)))]
   5.862 +  | mv_mul(x::p1,[],_) = [(0,mv_null2(#2(x)))] 
   5.863 +  | mv_mul(x::p1,y::p2,order) = mv_shorten(rev(sort (mv_geq order) (mv_mmul(x,y) :: (mv_mul(p1,y::p2,order) @
   5.864 +									    mv_mul([x],p2,order)))),order);
   5.865 +
   5.866 +(*. gets the maximum value of a list .*)
   5.867 +fun mv_getmax([])=0
   5.868 +  | mv_getmax(x::p1)= let 
   5.869 +		       val m=mv_getmax(p1);
   5.870 +		   in
   5.871 +		       if m>x then m
   5.872 +		       else x
   5.873 +		   end;
   5.874 +(*. calculates the maximum degree of an multivariate polynomial .*)
   5.875 +fun mv_grad([]:mv_poly) = 0 
   5.876 +  | mv_grad(p1:mv_poly)= mv_getmax((map mv_addlist) ((map #2) p1));
   5.877 +
   5.878 +(*. converts the sign of a value .*)
   5.879 +fun mv_minus(x)=(~1) * x;
   5.880 +
   5.881 +(*. converts the sign of all coefficients of a polynomial .*)
   5.882 +fun mv_minus2([]:mv_poly)=[]:mv_poly
   5.883 +  | mv_minus2(p1)=(mv_minus(#1(hd(p1))),#2(hd(p1)))::(mv_minus2(tl(p1)));
   5.884 +
   5.885 +(*. searches for a negativ value in a list .*)
   5.886 +fun mv_is_negativ([])=false
   5.887 +  | mv_is_negativ(x::xs)=if x<0 then true else mv_is_negativ(xs);
   5.888 +
   5.889 +(*. division of monomials .*)
   5.890 +fun mv_mdiv((0,[]):mv_monom,_:mv_monom)=(0,[]):mv_monom
   5.891 +  | mv_mdiv(_,(0,[]))= raise error ("RATIONALS_MV_MDIV_EXCEPTION Division by 0 ")
   5.892 +  | mv_mdiv(p1:mv_monom,p2:mv_monom)= 
   5.893 +    let
   5.894 +	val c=ref (#1(p2));
   5.895 +	val pp=ref [];
   5.896 +    in 
   5.897 +	(
   5.898 +	 if !c=0 then raise error("MV_MDIV_EXCEPTION Dividing by zero")
   5.899 +	 else c:=(#1(p1) div #1(p2));
   5.900 +	     if #1(p2)<>0 then 
   5.901 +		 (
   5.902 +		  pp:=(#2(mv_mmul((1,#2(p1)),(1,(map mv_minus) (#2(p2))))));
   5.903 +		  if mv_is_negativ(!pp) then (0,!pp)
   5.904 +		  else (!c,!pp) 
   5.905 +		      )
   5.906 +	     else raise error("MV_MDIV_EXCEPTION Dividing by empty Polynom")
   5.907 +		 )
   5.908 +    end;
   5.909 +
   5.910 +(*. prints a polynom for (internal use only) .*)
   5.911 +fun mv_print_poly([]:mv_poly)=print("[]\n")
   5.912 +  | mv_print_poly((x,y)::[])= print("("^BasisLibrary.Int.toString(x)^","^ints2str(y)^")\n")
   5.913 +  | mv_print_poly((x,y)::p1) = (print("("^BasisLibrary.Int.toString(x)^","^ints2str(y)^"),");mv_print_poly(p1));
   5.914 +
   5.915 +
   5.916 +(*. division of two multivariate polynomials .*) 
   5.917 +fun mv_division([]:mv_poly,g:mv_poly,order)=([]:mv_poly,[]:mv_poly)
   5.918 +  | mv_division(f,[],order)= raise error ("RATIONALS_MV_DIVISION_EXCEPTION Division by zero")
   5.919 +  | mv_division(f,g,order)=
   5.920 +    let 
   5.921 +	val r=ref [];
   5.922 +	val q=ref [];
   5.923 +	val g'=ref [];
   5.924 +	val k=ref 0;
   5.925 +	val m=ref (0,[0]);
   5.926 +	val exit=ref 0;
   5.927 +    in
   5.928 +	r := rev(sort (mv_geq order) (mv_shorten(f,order)));
   5.929 +	g':= rev(sort (mv_geq order) (mv_shorten(g,order)));
   5.930 +	if #1(hd(!g'))=0 then raise error("RATIONALS_MV_DIVISION_EXCEPTION: Dividing by zero") else ();
   5.931 +	if  (mv_monom_greater (hd(!g'),hd(!r),order)) then ([(0,mv_null2(#2(hd(f))))],(!r))
   5.932 +	else
   5.933 +	    (
   5.934 +	     exit:=0;
   5.935 +	     while (if (!exit)=0 then not(mv_monom_greater (hd(!g'),hd(!r),order)) else false) do
   5.936 +		 (
   5.937 +		  if (#1(mv_lm(!g',order)))<>0 then m:=mv_mdiv(mv_lm(!r,order),mv_lm(!g',order))
   5.938 +		  else raise error ("RATIONALS_MV_DIVISION_EXCEPTION: Dividing by zero");	  
   5.939 +		  if #1(!m)<>0 then
   5.940 +		      ( 
   5.941 +		       q:=(!m)::(!q);
   5.942 +		       r:=mv_add((!r),mv_minus2(mv_mul(!g',[!m],order)),order)
   5.943 +		       )
   5.944 +		  else exit:=1;
   5.945 +		  if (if length(!r)<>0 then length(!g')<>0 else false) then ()
   5.946 +		  else (exit:=1)
   5.947 +		  );
   5.948 +		 (rev(!q),!r)
   5.949 +		 )
   5.950 +    end;
   5.951 +
   5.952 +(*. multiplies a polynomial with an integer .*)
   5.953 +fun mv_skalar_mul([]:mv_poly,c) = []:mv_poly
   5.954 +  | mv_skalar_mul((x,y)::p1,c) = ((x * c),y)::mv_skalar_mul(p1,c); 
   5.955 +
   5.956 +(*. inserts the a first variable into an polynomial with exponent v .*)
   5.957 +fun mv_correct([]:mv_poly,v:int)=[]:mv_poly
   5.958 +  | mv_correct((x,y)::list,v:int)=(x,v::y)::mv_correct(list,v);
   5.959 +
   5.960 +(*. multivariate case .*)
   5.961 +
   5.962 +(*. decides if x is a factor of y .*)
   5.963 +fun mv_divides([]:mv_poly,[]:mv_poly)=  raise error("RATIONALS_MV_DIVIDES_EXCEPTION: division by zero")
   5.964 +  | mv_divides(x,[]) =  raise error("RATIONALS_MV_DIVIDES_EXCEPTION: division by zero")
   5.965 +  | mv_divides(x:mv_poly,y:mv_poly) = #2(mv_division(y,x,LEX_))=[];
   5.966 +
   5.967 +(*. gets the maximum of a and b .*)
   5.968 +fun mv_max(a,b) = if a>b then a else b;
   5.969 +
   5.970 +(*. gets the maximum exponent of a mv polynomial in the lexicographic term order .*)
   5.971 +fun mv_deg([]:mv_poly) = 0  
   5.972 +  | mv_deg(p1)=
   5.973 +    let
   5.974 +	val p1'=mv_shorten(p1,LEX_);
   5.975 +    in
   5.976 +	if length(p1')=0 then 0 
   5.977 +	else mv_max(hd(#2(hd(p1'))),mv_deg(tl(p1')))
   5.978 +    end;
   5.979 +
   5.980 +(*. gets the maximum exponent of a mv polynomial in the reverse lexicographic term order .*)
   5.981 +fun mv_deg2([]:mv_poly) = 0
   5.982 +  | mv_deg2(p1)=
   5.983 +    let
   5.984 +	val p1'=mv_shorten(p1,LEX_);
   5.985 +    in
   5.986 +	if length(p1')=0 then 0 
   5.987 +	else mv_max(hd(rev(#2(hd(p1')))),mv_deg2(tl(p1')))
   5.988 +    end;
   5.989 +
   5.990 +(*. evaluates the mv polynomial at the value v of the main variable .*)
   5.991 +fun mv_subs([]:mv_poly,v) = []:mv_poly
   5.992 +  | mv_subs((c,e)::p1:mv_poly,v) = mv_skalar_mul(mv_cut([(c,e)]),power v (hd(e))) @ mv_subs(p1,v);
   5.993 +
   5.994 +(*. calculates the content of a uv-polynomial in mv-representation .*)
   5.995 +fun uv_content2([]:mv_poly) = 0
   5.996 +  | uv_content2((c,e)::p1) = (gcd_int c (uv_content2(p1)));
   5.997 +
   5.998 +(*. converts a uv-polynomial from mv-representation to  uv-representation .*)
   5.999 +fun uv_to_list ([]:mv_poly)=[]:uv_poly
  5.1000 +  | uv_to_list ((c1,e1)::others) = 
  5.1001 +    let
  5.1002 +	val count=ref 0;
  5.1003 +	val max=mv_grad((c1,e1)::others); 
  5.1004 +	val help=ref ((c1,e1)::others);
  5.1005 +	val list=ref [];
  5.1006 +    in
  5.1007 +	if length(e1)>1 then raise error ("RATIONALS_TO_LIST_EXCEPTION: not univariate")
  5.1008 +	else if length(e1)=0 then [c1]
  5.1009 +	     else
  5.1010 +		 (
  5.1011 +		  count:=0;
  5.1012 +		  while (!count)<=max do
  5.1013 +		      (
  5.1014 +		       if length(!help)>0 andalso hd(#2(hd(!help)))=max-(!count) then 
  5.1015 +			   (
  5.1016 +			    list:=(#1(hd(!help)))::(!list);		       
  5.1017 +			    help:=tl(!help) 
  5.1018 +			    )
  5.1019 +		       else 
  5.1020 +			   (
  5.1021 +			    list:= 0::(!list)
  5.1022 +			    );
  5.1023 +		       count := (!count) + 1
  5.1024 +		       );
  5.1025 +		      (!list)
  5.1026 +		      )
  5.1027 +    end;
  5.1028 +
  5.1029 +(*. converts a uv-polynomial from uv-representation to mv-representation .*)
  5.1030 +fun uv_to_poly ([]:uv_poly) = []:mv_poly
  5.1031 +  | uv_to_poly p1 = 
  5.1032 +    let
  5.1033 +	val count=ref 0;
  5.1034 +	val help=ref p1;
  5.1035 +	val list=ref [];
  5.1036 +    in
  5.1037 +	while length(!help)>0 do
  5.1038 +	    (
  5.1039 +	     if hd(!help)=0 then ()
  5.1040 +	     else list:=(hd(!help),[!count])::(!list);
  5.1041 +	     count:=(!count)+1;
  5.1042 +	     help:=tl(!help)
  5.1043 +	     );
  5.1044 +	    (!list)
  5.1045 +    end;
  5.1046 +
  5.1047 +(*. univariate gcd calculation from polynomials in multivariate representation .*)
  5.1048 +fun uv_gcd ([]:mv_poly) p2 = p2
  5.1049 +  | uv_gcd p1 ([]:mv_poly) = p1
  5.1050 +  | uv_gcd p1 [(c,[e])] = 
  5.1051 +    let 
  5.1052 +	val list=ref (rev(sort (mv_geq LEX_) (mv_shorten(p1,LEX_))));
  5.1053 +	val min=uv_mod_min(e,(hd(#2(hd(rev(!list))))));
  5.1054 +    in
  5.1055 +	[(gcd_int (uv_content2(p1)) c,[min])]
  5.1056 +    end
  5.1057 +  | uv_gcd [(c,[e])] p2 = 
  5.1058 +    let 
  5.1059 +	val list=ref (rev(sort (mv_geq LEX_) (mv_shorten(p2,LEX_))));
  5.1060 +	val min=uv_mod_min(e,(hd(#2(hd(rev(!list))))));
  5.1061 +    in
  5.1062 +	[(gcd_int (uv_content2(p2)) c,[min])]
  5.1063 +    end 
  5.1064 +  | uv_gcd p11 p22 = uv_to_poly(uv_mod_gcd (uv_to_list(mv_shorten(p11,LEX_))) (uv_to_list(mv_shorten(p22,LEX_))));
  5.1065 +
  5.1066 +(*. help function for the newton interpolation .*)
  5.1067 +fun mv_newton_help ([]:mv_poly list,k:int) = []:mv_poly list
  5.1068 +  | mv_newton_help (pl:mv_poly list,k) = 
  5.1069 +    let
  5.1070 +	val x=ref (rev(pl));
  5.1071 +	val t=ref [];
  5.1072 +	val y=ref [];
  5.1073 +	val n=ref 1;
  5.1074 +	val n1=ref[];
  5.1075 +    in
  5.1076 +	(
  5.1077 +	 while length(!x)>1 do 
  5.1078 +	     (
  5.1079 +	      if length(hd(!x))>0 then n1:=mv_null2(#2(hd(hd(!x))))
  5.1080 +	      else if length(hd(tl(!x)))>0 then n1:=mv_null2(#2(hd(hd(tl(!x)))))
  5.1081 +		   else n1:=[]; 
  5.1082 +	      t:= #1(mv_division(mv_add(hd(!x),mv_skalar_mul(hd(tl(!x)),~1),LEX_),[(k,!n1)],LEX_)); 
  5.1083 +	      y:=(!t)::(!y);
  5.1084 +	      x:=tl(!x)
  5.1085 +	      );
  5.1086 +	 (!y)
  5.1087 +	 )
  5.1088 +    end;
  5.1089 +    
  5.1090 +(*. help function for the newton interpolation .*)
  5.1091 +fun mv_newton_add ([]:mv_poly list) t = []:mv_poly
  5.1092 +  | mv_newton_add [x:mv_poly] t = x 
  5.1093 +  | mv_newton_add (pl:mv_poly list) t = 
  5.1094 +    let
  5.1095 +	val expos=ref [];
  5.1096 +	val pll=ref pl;
  5.1097 +    in
  5.1098 +	(
  5.1099 +
  5.1100 +	 while length(!pll)>0 andalso hd(!pll)=[]  do 
  5.1101 +	     ( 
  5.1102 +	      pll:=tl(!pll)
  5.1103 +	      ); 
  5.1104 +	 if length(!pll)>0 then expos:= #2(hd(hd(!pll))) else expos:=[];
  5.1105 +	 mv_add(hd(pl),
  5.1106 +		mv_mul(
  5.1107 +		       mv_add(mv_correct(mv_cut([(1,mv_null2(!expos))]),1),[(~t,mv_null2(!expos))],LEX_),
  5.1108 +		       mv_newton_add (tl(pl)) (t+1),
  5.1109 +		       LEX_
  5.1110 +		       ),
  5.1111 +		LEX_)
  5.1112 +	 )
  5.1113 +    end;
  5.1114 +
  5.1115 +(*. calculates the newton interpolation with polynomial coefficients .*)
  5.1116 +(*. step-depth is 1 and if the result is not an integerpolynomial .*)
  5.1117 +(*. this function returns [] .*)
  5.1118 +fun mv_newton ([]:(mv_poly) list) = []:mv_poly 
  5.1119 +  | mv_newton ([mp]:(mv_poly) list) = mp:mv_poly
  5.1120 +  | mv_newton pl =
  5.1121 +    let
  5.1122 +	val c=ref pl;
  5.1123 +	val c1=ref [];
  5.1124 +	val n=length(pl);
  5.1125 +	val k=ref 1;
  5.1126 +	val i=ref n;
  5.1127 +	val ppl=ref [];
  5.1128 +    in
  5.1129 +	c1:=hd(pl)::[];
  5.1130 +	c:=mv_newton_help(!c,!k);
  5.1131 +	c1:=(hd(!c))::(!c1);
  5.1132 +	while(length(!c)>1 andalso !k<n) do
  5.1133 +	    (	 
  5.1134 +	     k:=(!k)+1; 
  5.1135 +	     while  length(!c)>0 andalso hd(!c)=[] do c:=tl(!c); 	  
  5.1136 +	     if !c=[] then () else c:=mv_newton_help(!c,!k);
  5.1137 +	     ppl:= !c;
  5.1138 +	     if !c=[] then () else  c1:=(hd(!c))::(!c1)
  5.1139 +	     );
  5.1140 +	while  hd(!c1)=[] do c1:=tl(!c1);
  5.1141 +	c1:=rev(!c1);
  5.1142 +	ppl:= !c1;
  5.1143 +	mv_newton_add (!c1) 1
  5.1144 +    end;
  5.1145 +
  5.1146 +(*. sets the exponents of the first variable to zero .*)
  5.1147 +fun mv_null3([]:mv_poly)    = []:mv_poly
  5.1148 +  | mv_null3((x,y)::xs) = (x,0::tl(y))::mv_null3(xs);
  5.1149 +
  5.1150 +(*. calculates the minimum exponents of a multivariate polynomial .*)
  5.1151 +fun mv_min_pp([]:mv_poly)=[]
  5.1152 +  | mv_min_pp((c,e)::xs)=
  5.1153 +    let
  5.1154 +	val y=ref xs;
  5.1155 +	val x=ref [];
  5.1156 +    in
  5.1157 +	(
  5.1158 +	 x:=e;
  5.1159 +	 while length(!y)>0 do
  5.1160 +	     (
  5.1161 +	      x:=(map uv_mod_min) ((!x) ~~ (#2(hd(!y))));
  5.1162 +	      y:=tl(!y)
  5.1163 +	      );
  5.1164 +	 !x
  5.1165 +	 )
  5.1166 +    end;
  5.1167 +
  5.1168 +(*. checks if all elements of the list have value zero .*)
  5.1169 +fun list_is_null [] = true 
  5.1170 +  | list_is_null (x::xs) = (x=0 andalso list_is_null(xs)); 
  5.1171 +
  5.1172 +(* check if main variable is zero*)
  5.1173 +fun main_zero (ms : mv_poly) = (list_is_null o (map (hd o #2))) ms;
  5.1174 +
  5.1175 +(*. calculates the content of an polynomial .*)
  5.1176 +fun mv_content([]:mv_poly) = []:mv_poly
  5.1177 +  | mv_content(p1) = 
  5.1178 +    let
  5.1179 +	val list=ref (rev(sort (mv_geq LEX_) (mv_shorten(p1,LEX_))));
  5.1180 +	val test=ref (hd(#2(hd(!list))));
  5.1181 +	val result=ref []; 
  5.1182 +	val min=(hd(#2(hd(rev(!list)))));
  5.1183 +    in
  5.1184 +	(
  5.1185 +	 if length(!list)>1 then
  5.1186 +	     (
  5.1187 +	      while (if length(!list)>0 then (hd(#2(hd(!list)))=(!test)) else false) do
  5.1188 +		  (
  5.1189 +		   result:=(#1(hd(!list)),tl(#2(hd(!list))))::(!result);
  5.1190 +		   
  5.1191 +		   if length(!list)<1 then list:=[]
  5.1192 +		   else list:=tl(!list) 
  5.1193 +		       
  5.1194 +		       );		  
  5.1195 +		  if length(!list)>0 then  
  5.1196 +		   ( 
  5.1197 +		    list:=mv_gcd (!result) (mv_cut(mv_content(!list))) 
  5.1198 +		    ) 
  5.1199 +		  else list:=(!result); 
  5.1200 +		  list:=mv_correct(!list,0);  
  5.1201 +		  (!list) 
  5.1202 +		  )
  5.1203 +	 else
  5.1204 +	     (
  5.1205 +	      mv_null3(!list) 
  5.1206 +	      )
  5.1207 +	     )
  5.1208 +    end
  5.1209 +
  5.1210 +(*. calculates the primitiv part of a polynomial .*)
  5.1211 +and mv_pp([]:mv_poly) = []:mv_poly
  5.1212 +  | mv_pp(p1) = let
  5.1213 +		    val cont=ref []; 
  5.1214 +		    val pp=ref[];
  5.1215 +		in
  5.1216 +		    cont:=mv_content(p1);
  5.1217 +		    pp:=(#1(mv_division(p1,!cont,LEX_)));
  5.1218 +		    if !pp=[] 
  5.1219 +			then raise error("RATIONALS_MV_PP_EXCEPTION: Invalid Content ")
  5.1220 +		    else (!pp)
  5.1221 +		end
  5.1222 +
  5.1223 +(*. calculates the gcd of two multivariate polynomials with a modular approach .*)
  5.1224 +and mv_gcd ([]:mv_poly) ([]:mv_poly) :mv_poly= []:mv_poly
  5.1225 +  | mv_gcd ([]:mv_poly) (p2) :mv_poly= p2:mv_poly
  5.1226 +  | mv_gcd (p1:mv_poly) ([]) :mv_poly= p1:mv_poly
  5.1227 +  | mv_gcd ([(x,xs)]:mv_poly) ([(y,ys)]):mv_poly = 
  5.1228 +     let
  5.1229 +      val xpoly:mv_poly = [(x,xs)];
  5.1230 +      val ypoly:mv_poly = [(y,ys)];
  5.1231 +     in 
  5.1232 +	(
  5.1233 +	 if xs=ys then [((gcd_int x y),xs)]
  5.1234 +	 else [((gcd_int x y),(map uv_mod_min)(xs~~ys))]:mv_poly
  5.1235 +        )
  5.1236 +    end 
  5.1237 +  | mv_gcd (p1:mv_poly) ([(y,ys)]) :mv_poly= 
  5.1238 +	(
  5.1239 +	 [(gcd_int (uv_content2(p1)) (y),(map  uv_mod_min)(mv_min_pp(p1)~~ys))]:mv_poly
  5.1240 +	)
  5.1241 +  | mv_gcd ([(y,ys)]:mv_poly) (p2):mv_poly = 
  5.1242 +	(
  5.1243 +         [(gcd_int (uv_content2(p2)) (y),(map  uv_mod_min)(mv_min_pp(p2)~~ys))]:mv_poly
  5.1244 +        )
  5.1245 +  | mv_gcd (p1':mv_poly) (p2':mv_poly):mv_poly=
  5.1246 +    let
  5.1247 +	val vc=length(#2(hd(p1')));
  5.1248 +	val cont = 
  5.1249 +		  (
  5.1250 +                   if main_zero(mv_content(p1')) andalso 
  5.1251 +                     (main_zero(mv_content(p2'))) then
  5.1252 +                     mv_correct((mv_gcd (mv_cut(mv_content(p1'))) (mv_cut(mv_content(p2')))),0)
  5.1253 +                   else 
  5.1254 +                     mv_gcd (mv_content(p1')) (mv_content(p2'))
  5.1255 +                  );
  5.1256 +	val p1= #1(mv_division(p1',mv_content(p1'),LEX_));
  5.1257 +	val p2= #1(mv_division(p2',mv_content(p2'),LEX_)); 
  5.1258 +	val gcd=ref [];
  5.1259 +	val candidate=ref [];
  5.1260 +	val interpolation_list=ref [];
  5.1261 +	val delta=ref [];
  5.1262 +        val p1r = ref [];
  5.1263 +        val p2r = ref [];
  5.1264 +        val p1r' = ref [];
  5.1265 +        val p2r' = ref [];
  5.1266 +	val factor=ref [];
  5.1267 +	val r=ref 0;
  5.1268 +	val gcd_r=ref [];
  5.1269 +	val d=ref 0;
  5.1270 +	val exit=ref 0;
  5.1271 +	val current_degree=ref 99999; (*. FIXME: unlimited ! .*)
  5.1272 +    in
  5.1273 +	(
  5.1274 +	 if vc<2 then (* areUnivariate(p1',p2') *)
  5.1275 +	     (
  5.1276 +	      gcd:=uv_gcd (mv_shorten(p1',LEX_)) (mv_shorten(p2',LEX_))
  5.1277 +	      )
  5.1278 +	 else
  5.1279 +	     (
  5.1280 +	      while !exit=0 do
  5.1281 +		  (
  5.1282 +		   r:=(!r)+1;
  5.1283 +                   p1r := mv_lc(p1,LEX_);
  5.1284 +		   p2r := mv_lc(p2,LEX_);
  5.1285 +                   if main_zero(!p1r) andalso
  5.1286 +                      main_zero(!p2r) 
  5.1287 +                   then
  5.1288 +                       (
  5.1289 +                        delta := mv_correct((mv_gcd (mv_cut (!p1r)) (mv_cut (!p2r))),0)
  5.1290 +                       )
  5.1291 +                   else
  5.1292 +                       (
  5.1293 +		        delta := mv_gcd (!p1r) (!p2r)
  5.1294 +                       );
  5.1295 +		   (*if mv_shorten(mv_subs(!p1r,!r),LEX_)=[] andalso 
  5.1296 +		      mv_shorten(mv_subs(!p2r,!r),LEX_)=[] *)
  5.1297 +		   if mv_lc2(mv_shorten(mv_subs(!p1r,!r),LEX_),LEX_)=0 andalso 
  5.1298 +		      mv_lc2(mv_shorten(mv_subs(!p2r,!r),LEX_),LEX_)=0 
  5.1299 +                   then 
  5.1300 +                       (
  5.1301 +		       )
  5.1302 +		   else 
  5.1303 +		       (
  5.1304 +			gcd_r:=mv_shorten(mv_gcd (mv_shorten(mv_subs(p1,!r),LEX_)) 
  5.1305 +					         (mv_shorten(mv_subs(p2,!r),LEX_)) ,LEX_);
  5.1306 +			gcd_r:= #1(mv_division(mv_mul(mv_correct(mv_subs(!delta,!r),0),!gcd_r,LEX_),
  5.1307 +					       mv_correct(mv_lc(!gcd_r,LEX_),0),LEX_));
  5.1308 +			d:=mv_deg2(!gcd_r); (* deg(gcd_r,z) *)
  5.1309 +			if (!d < !current_degree) then 
  5.1310 +			    (
  5.1311 +			     current_degree:= !d;
  5.1312 +			     interpolation_list:=mv_correct(!gcd_r,0)::(!interpolation_list)
  5.1313 +			     )
  5.1314 +			else
  5.1315 +			    (
  5.1316 +			     if (!d = !current_degree) then
  5.1317 +				 (
  5.1318 +				  interpolation_list:=mv_correct(!gcd_r,0)::(!interpolation_list)
  5.1319 +				  )
  5.1320 +			     else () 
  5.1321 +				 )
  5.1322 +			    );
  5.1323 +		      if length(!interpolation_list)> uv_mod_min(mv_deg(p1),mv_deg(p2)) then 
  5.1324 +			  (
  5.1325 +			   candidate := mv_newton(rev(!interpolation_list));
  5.1326 +			   if !candidate=[] then ()
  5.1327 +			   else
  5.1328 +			       (
  5.1329 +				candidate:=mv_pp(!candidate);
  5.1330 +				if mv_divides(!candidate,p1) andalso mv_divides(!candidate,p2) then
  5.1331 +				    (
  5.1332 +				     gcd:= mv_mul(!candidate,cont,LEX_);
  5.1333 +				     exit:=1
  5.1334 +				     )
  5.1335 +				else ()
  5.1336 +				    );
  5.1337 +			       interpolation_list:=[mv_correct(!gcd_r,0)]
  5.1338 +			       )
  5.1339 +		      else ()
  5.1340 +			  )
  5.1341 +	     );
  5.1342 +	     (!gcd):mv_poly
  5.1343 +	     )
  5.1344 +    end;	
  5.1345 +
  5.1346 +
  5.1347 +(*. calculates the least common divisor of two polynomials .*)
  5.1348 +fun mv_lcm (p1:mv_poly) (p2:mv_poly) :mv_poly = 
  5.1349 +    (
  5.1350 +     #1(mv_division(mv_mul(p1,p2,LEX_),mv_gcd p1 p2,LEX_))
  5.1351 +     );
  5.1352 +
  5.1353 +(*. gets the variables (strings) of a term .*)
  5.1354 +fun get_vars(term1) = (map free2str) (vars term1); (*["a","b","c"]; *)
  5.1355 +
  5.1356 +(*. counts the negative coefficents in a polynomial .*)
  5.1357 +fun count_neg ([]:mv_poly) = 0 
  5.1358 +  | count_neg ((c,e)::xs) = if c<0 then 1+count_neg xs
  5.1359 +			  else count_neg xs;
  5.1360 +
  5.1361 +(*. help function for is_polynomial  
  5.1362 +    checks the order of the operators .*)
  5.1363 +fun test_polynomial (Const ("uminus",_) $ Free (str,_)) _ = true (*WN.13.3.03*)
  5.1364 +  | test_polynomial (t as Free(str,_)) v = true
  5.1365 +  | test_polynomial (t as Const ("op *",_) $ t1 $ t2) v = if v="^" then false
  5.1366 +						     else (test_polynomial t1 "*") andalso (test_polynomial t2 "*")
  5.1367 +  | test_polynomial (t as Const ("op +",_) $ t1 $ t2) v = if v="*" orelse v="^" then false 
  5.1368 +							  else (test_polynomial t1 " ") andalso (test_polynomial t2 " ")
  5.1369 +  | test_polynomial (t as Const ("Atools.pow",_) $ t1 $ t2) v = (test_polynomial t1 "^") andalso (test_polynomial t2 "^")
  5.1370 +  | test_polynomial _ v = false;  
  5.1371 +
  5.1372 +(*. tests if a term is a polynomial .*)  
  5.1373 +fun is_polynomial t = test_polynomial t " ";
  5.1374 +
  5.1375 +(*. help function for is_expanded 
  5.1376 +    checks the order of the operators .*)
  5.1377 +fun test_exp (t as Free(str,_)) v = true 
  5.1378 +  | test_exp (t as Const ("op *",_) $ t1 $ t2) v = if v="^" then false
  5.1379 +						     else (test_exp t1 "*") andalso (test_exp t2 "*")
  5.1380 +  | test_exp (t as Const ("op +",_) $ t1 $ t2) v = if v="*" orelse v="^" then false 
  5.1381 +							  else (test_exp t1 " ") andalso (test_exp t2 " ") 
  5.1382 +  | test_exp (t as Const ("op -",_) $ t1 $ t2) v = if v="*" orelse v="^" then false 
  5.1383 +							  else (test_exp t1 " ") andalso (test_exp t2 " ")
  5.1384 +  | test_exp (t as Const ("Atools.pow",_) $ t1 $ t2) v = (test_exp t1 "^") andalso (test_exp t2 "^")
  5.1385 +  | test_exp  _ v = false;
  5.1386 +
  5.1387 +
  5.1388 +(*. help function for check_coeff: 
  5.1389 +    converts the term to a list of coefficients .*) 
  5.1390 +fun term2coef' (t as Free(str,_(*typ*))) v :mv_poly option = 
  5.1391 +    let
  5.1392 +	val x=ref NONE;
  5.1393 +	val len=ref 0;
  5.1394 +	val vl=ref [];
  5.1395 +	val vh=ref [];
  5.1396 +	val i=ref 0;
  5.1397 +    in 
  5.1398 +	if is_numeral str then
  5.1399 +	    (
  5.1400 +	     SOME [(((the o int_of_str) str),mv_null2(v))] handle _ => NONE
  5.1401 +		 )
  5.1402 +	else (* variable *)
  5.1403 +	    (
  5.1404 +	     len:=length(v);
  5.1405 +	     vh:=v;
  5.1406 +	     while ((!len)>(!i)) do
  5.1407 +		 (
  5.1408 +		  if str=hd((!vh)) then
  5.1409 +		      (
  5.1410 +		       vl:=1::(!vl)
  5.1411 +		       )
  5.1412 +		  else 
  5.1413 +		      (
  5.1414 +		       vl:=0::(!vl)
  5.1415 +		       );
  5.1416 +		      vh:=tl(!vh);
  5.1417 +		      i:=(!i)+1    
  5.1418 +		      );		
  5.1419 +		 SOME [(1,rev(!vl))] handle _ => NONE
  5.1420 +	    )
  5.1421 +    end
  5.1422 +  | term2coef' (Const ("op *",_) $ t1 $ t2) v :mv_poly option= 
  5.1423 +    let
  5.1424 +	val t1pp=ref [];
  5.1425 +	val t2pp=ref [];
  5.1426 +	val t1c=ref 0;
  5.1427 +	val t2c=ref 0;
  5.1428 +    in
  5.1429 +	(
  5.1430 +	 t1pp:=(#2(hd(the(term2coef' t1 v))));
  5.1431 +	 t2pp:=(#2(hd(the(term2coef' t2 v))));
  5.1432 +	 t1c:=(#1(hd(the(term2coef' t1 v))));
  5.1433 +	 t2c:=(#1(hd(the(term2coef' t2 v))));
  5.1434 +	
  5.1435 +	 SOME [( (!t1c)*(!t2c) ,( (map op+) ((!t1pp)~~(!t2pp)) ) )] handle _ => NONE 
  5.1436 +		
  5.1437 +	 )
  5.1438 +    end
  5.1439 +  | term2coef' (Const ("Atools.pow",_) $ (t1 as Free (str1,_)) $ (t2 as Free (str2,_))) v :mv_poly option= 
  5.1440 +    let
  5.1441 +	val x=ref NONE;
  5.1442 +	val len=ref 0;
  5.1443 +	val vl=ref [];
  5.1444 +	val vh=ref [];
  5.1445 +	val vtemp=ref [];
  5.1446 +	val i=ref 0;	 
  5.1447 +    in
  5.1448 +    (
  5.1449 +     if (not o is_numeral) str1 andalso is_numeral str2 then
  5.1450 +	 (
  5.1451 +	  len:=length(v);
  5.1452 +	  vh:=v;
  5.1453 +
  5.1454 +	  while ((!len)>(!i)) do
  5.1455 +	      (
  5.1456 +	       if str1=hd((!vh)) then
  5.1457 +		   (
  5.1458 +		    vl:=((the o int_of_str) str2)::(!vl)
  5.1459 +		    )
  5.1460 +	       else 
  5.1461 +		   (
  5.1462 +		    vl:=0::(!vl)
  5.1463 +		    );
  5.1464 +		   vh:=tl(!vh);
  5.1465 +		   i:=(!i)+1     
  5.1466 +		   );
  5.1467 +	      SOME [(1,rev(!vl))] handle _ => NONE
  5.1468 +	      )
  5.1469 +     else raise error ("RATIONALS_TERM2COEF_EXCEPTION 1: Invalid term")
  5.1470 +	 )
  5.1471 +    end
  5.1472 +  | term2coef' (Const ("op +",_) $ t1 $ t2) v :mv_poly option= 
  5.1473 +    (
  5.1474 +     SOME ((the(term2coef' t1 v)) @ (the(term2coef' t2 v))) handle _ => NONE
  5.1475 +	 )
  5.1476 +  | term2coef' (Const ("op -",_) $ t1 $ t2) v :mv_poly option= 
  5.1477 +    (
  5.1478 +     SOME ((the(term2coef' t1 v)) @ mv_skalar_mul((the(term2coef' t2 v)),1)) handle _ => NONE
  5.1479 +	 )
  5.1480 +  | term2coef' (term) v = raise error ("RATIONALS_TERM2COEF_EXCEPTION 2: Invalid term");
  5.1481 +
  5.1482 +(*. checks if all coefficients of a polynomial are positiv (except the first) .*)
  5.1483 +fun check_coeff t = (* erste Koeffizient kann <0 sein !!! *)
  5.1484 +    if count_neg(tl(the(term2coef' t (get_vars(t)))))=0 then true 
  5.1485 +    else false;
  5.1486 +
  5.1487 +(*. checks for expanded term [3] .*)
  5.1488 +fun is_expanded t = test_exp t " " andalso check_coeff(t); 
  5.1489 +
  5.1490 +(*WN.7.3.03 Hilfsfunktion f"ur term2poly'*)
  5.1491 +fun mk_monom v' p vs = 
  5.1492 +    let fun conv p (v: string) = if v'= v then p else 0
  5.1493 +    in map (conv p) vs end;
  5.1494 +(* mk_monom "y" 5 ["a","b","x","y","z"];
  5.1495 +val it = [0,0,0,5,0] : int list*)
  5.1496 +
  5.1497 +(*. this function converts the term representation into the internal representation mv_poly .*)
  5.1498 +fun term2poly' (Const ("uminus",_) $ Free (str,_)) v = (*WN.7.3.03*)
  5.1499 +    if is_numeral str 
  5.1500 +    then SOME [((the o int_of_str) ("-"^str), mk_monom "#" 0 v)]
  5.1501 +    else SOME [(~1, mk_monom str 1 v)]
  5.1502 +
  5.1503 +  | term2poly' (Free(str,_)) v :mv_poly option = 
  5.1504 +    let
  5.1505 +	val x=ref NONE;
  5.1506 +	val len=ref 0;
  5.1507 +	val vl=ref [];
  5.1508 +	val vh=ref [];
  5.1509 +	val i=ref 0;
  5.1510 +    in 
  5.1511 +	if is_numeral str then
  5.1512 +	    (
  5.1513 +	     SOME [(((the o int_of_str) str),mv_null2 v)] handle _ => NONE
  5.1514 +		 )
  5.1515 +	else (* variable *)
  5.1516 +	    (
  5.1517 +	     len:=length v;
  5.1518 +	     vh:= v;
  5.1519 +	     while ((!len)>(!i)) do
  5.1520 +		 (
  5.1521 +		  if str=hd((!vh)) then
  5.1522 +		      (
  5.1523 +		       vl:=1::(!vl)
  5.1524 +		       )
  5.1525 +		  else 
  5.1526 +		      (
  5.1527 +		       vl:=0::(!vl)
  5.1528 +		       );
  5.1529 +		      vh:=tl(!vh);
  5.1530 +		      i:=(!i)+1    
  5.1531 +		      );		
  5.1532 +		 SOME [(1,rev(!vl))] handle _ => NONE
  5.1533 +	    )
  5.1534 +    end
  5.1535 +  | term2poly' (Const ("op *",_) $ t1 $ t2) v :mv_poly option= 
  5.1536 +    let
  5.1537 +	val t1pp=ref [];
  5.1538 +	val t2pp=ref [];
  5.1539 +	val t1c=ref 0;
  5.1540 +	val t2c=ref 0;
  5.1541 +    in
  5.1542 +	(
  5.1543 +	 t1pp:=(#2(hd(the(term2poly' t1 v))));
  5.1544 +	 t2pp:=(#2(hd(the(term2poly' t2 v))));
  5.1545 +	 t1c:=(#1(hd(the(term2poly' t1 v))));
  5.1546 +	 t2c:=(#1(hd(the(term2poly' t2 v))));
  5.1547 +	
  5.1548 +	 SOME [( (!t1c)*(!t2c) ,( (map op+) ((!t1pp)~~(!t2pp)) ) )] 
  5.1549 +	 handle _ => NONE 
  5.1550 +		
  5.1551 +	 )
  5.1552 +    end
  5.1553 +  | term2poly' (Const ("Atools.pow",_) $ (t1 as Free (str1,_)) $ 
  5.1554 +		      (t2 as Free (str2,_))) v :mv_poly option= 
  5.1555 +    let
  5.1556 +	val x=ref NONE;
  5.1557 +	val len=ref 0;
  5.1558 +	val vl=ref [];
  5.1559 +	val vh=ref [];
  5.1560 +	val vtemp=ref [];
  5.1561 +	val i=ref 0;	 
  5.1562 +    in
  5.1563 +    (
  5.1564 +     if (not o is_numeral) str1 andalso is_numeral str2 then
  5.1565 +	 (
  5.1566 +	  len:=length(v);
  5.1567 +	  vh:=v;
  5.1568 +
  5.1569 +	  while ((!len)>(!i)) do
  5.1570 +	      (
  5.1571 +	       if str1=hd((!vh)) then
  5.1572 +		   (
  5.1573 +		    vl:=((the o int_of_str) str2)::(!vl)
  5.1574 +		    )
  5.1575 +	       else 
  5.1576 +		   (
  5.1577 +		    vl:=0::(!vl)
  5.1578 +		    );
  5.1579 +		   vh:=tl(!vh);
  5.1580 +		   i:=(!i)+1     
  5.1581 +		   );
  5.1582 +	      SOME [(1,rev(!vl))] handle _ => NONE
  5.1583 +	      )
  5.1584 +     else raise error ("RATIONALS_TERM2POLY_EXCEPTION 1: Invalid term")
  5.1585 +	 )
  5.1586 +    end
  5.1587 +  | term2poly' (Const ("op +",_) $ t1 $ t2) v :mv_poly option = 
  5.1588 +    (
  5.1589 +     SOME ((the(term2poly' t1 v)) @ (the(term2poly' t2 v))) handle _ => NONE
  5.1590 +	 )
  5.1591 +  | term2poly' (Const ("op -",_) $ t1 $ t2) v :mv_poly option = 
  5.1592 +    (
  5.1593 +     SOME ((the(term2poly' t1 v)) @ mv_skalar_mul((the(term2poly' t2 v)),~1)) handle _ => NONE
  5.1594 +	 )
  5.1595 +  | term2poly' (term) v = raise error ("RATIONALS_TERM2POLY_EXCEPTION 2: Invalid term");
  5.1596 +
  5.1597 +(*. translates an Isabelle term into internal representation.
  5.1598 +    term2poly
  5.1599 +    fn : term ->              (*normalform [2]                    *)
  5.1600 +    	 string list ->       (*for ...!!! BITTE DIE ERKLÄRUNG, 
  5.1601 +    			       DIE DU MIR LETZTES MAL GEGEBEN HAST*)
  5.1602 +    	 mv_monom list        (*internal representation           *)
  5.1603 +    		  option      (*the translation may fail with NONE*)
  5.1604 +.*)
  5.1605 +fun term2poly (t:term) v = 
  5.1606 +     if is_polynomial t then term2poly' t v
  5.1607 +     else raise error ("term2poly: invalid = "^(term2str t));
  5.1608 +
  5.1609 +(*. same as term2poly with automatic detection of the variables .*)
  5.1610 +fun term2polyx t = term2poly t (((map free2str) o vars) t); 
  5.1611 +
  5.1612 +(*. checks if the term is in expanded polynomial form and converts it into the internal representation .*)
  5.1613 +fun expanded2poly (t:term) v = 
  5.1614 +    (*if is_expanded t then*) term2poly' t v
  5.1615 +    (*else raise error ("RATIONALS_EXPANDED2POLY_EXCEPTION: Invalid Polynomial")*);
  5.1616 +
  5.1617 +(*. same as expanded2poly with automatic detection of the variables .*)
  5.1618 +fun expanded2polyx t = expanded2poly t (((map free2str) o vars) t);
  5.1619 +
  5.1620 +(*. converts a powerproduct into term representation .*)
  5.1621 +fun powerproduct2term(xs,v) =  
  5.1622 +    let
  5.1623 +	val xss=ref xs;
  5.1624 +	val vv=ref v;
  5.1625 +    in
  5.1626 +	(
  5.1627 +	 while hd(!xss)=0 do 
  5.1628 +	     (
  5.1629 +	      xss:=tl(!xss);
  5.1630 +	      vv:=tl(!vv)
  5.1631 +	      );
  5.1632 +	     
  5.1633 +	 if list_is_null(tl(!xss)) then 
  5.1634 +	     (
  5.1635 +	      if hd(!xss)=1 then Free(hd(!vv), HOLogic.realT)
  5.1636 +	      else
  5.1637 +		  (
  5.1638 +		   Const("Atools.pow",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1639 +		   Free(hd(!vv), HOLogic.realT) $  Free(str_of_int (hd(!xss)),HOLogic.realT)
  5.1640 +		   )
  5.1641 +	      )
  5.1642 +	 else
  5.1643 +	     (
  5.1644 +	      if hd(!xss)=1 then 
  5.1645 +		  ( 
  5.1646 +		   Const("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1647 +		   Free(hd(!vv), HOLogic.realT) $
  5.1648 +		   powerproduct2term(tl(!xss),tl(!vv))
  5.1649 +		   )
  5.1650 +	      else
  5.1651 +		  (
  5.1652 +		   Const("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1653 +		   (
  5.1654 +		    Const("Atools.pow",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1655 +		    Free(hd(!vv), HOLogic.realT) $  Free(str_of_int (hd(!xss)),HOLogic.realT)
  5.1656 +		    ) $
  5.1657 +		    powerproduct2term(tl(!xss),tl(!vv))
  5.1658 +		   )
  5.1659 +	      )
  5.1660 +	 )
  5.1661 +    end;
  5.1662 +
  5.1663 +(*. converts a monom into term representation .*)
  5.1664 +(*fun monom2term ((c,e):mv_monom, v:string list) = 
  5.1665 +    if c=0 then Free(str_of_int 0,HOLogic.realT)  
  5.1666 +    else
  5.1667 +	(
  5.1668 +	 if list_is_null(e) then
  5.1669 +	     ( 
  5.1670 +	      Free(str_of_int c,HOLogic.realT)  
  5.1671 +	      )
  5.1672 +	 else
  5.1673 +	     (
  5.1674 +	      if c=1 then 
  5.1675 +		  (
  5.1676 +		   powerproduct2term(e,v)
  5.1677 +		   )
  5.1678 +	      else
  5.1679 +		  (
  5.1680 +		   Const("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $
  5.1681 +		   Free(str_of_int c,HOLogic.realT)  $
  5.1682 +		   powerproduct2term(e,v)
  5.1683 +		   )
  5.1684 +		  )
  5.1685 +	     );*)
  5.1686 +
  5.1687 +
  5.1688 +(*fun monom2term ((i, is):mv_monom, v) = 
  5.1689 +    if list_is_null is 
  5.1690 +    then 
  5.1691 +	if i >= 0 
  5.1692 +	then Free (str_of_int i, HOLogic.realT)
  5.1693 +	else Const ("uminus", HOLogic.realT --> HOLogic.realT) $
  5.1694 +		   Free ((str_of_int o abs) i, HOLogic.realT)
  5.1695 +    else
  5.1696 +	if i > 0 
  5.1697 +	then Const ("op *", [HOLogic.realT,HOLogic.realT]---> HOLogic.realT) $
  5.1698 +		   (Free (str_of_int i, HOLogic.realT)) $
  5.1699 +		   powerproduct2term(is, v)
  5.1700 +	else Const ("op *", [HOLogic.realT,HOLogic.realT]---> HOLogic.realT) $
  5.1701 +		   (Const ("uminus", HOLogic.realT --> HOLogic.realT) $
  5.1702 +		   Free ((str_of_int o abs) i, HOLogic.realT)) $
  5.1703 +		   powerproduct2term(is, vs);---------------------------*)
  5.1704 +fun monom2term ((i, is) : mv_monom, vs) = 
  5.1705 +    if list_is_null is 
  5.1706 +    then Free (str_of_int i, HOLogic.realT)
  5.1707 +    else if i = 1
  5.1708 +    then powerproduct2term (is, vs)
  5.1709 +    else Const ("op *", [HOLogic.realT, HOLogic.realT] ---> HOLogic.realT) $
  5.1710 +	       (Free (str_of_int i, HOLogic.realT)) $
  5.1711 +	       powerproduct2term (is, vs);
  5.1712 +    
  5.1713 +(*. converts the internal polynomial representation into an Isabelle term.*)
  5.1714 +fun poly2term' ([] : mv_poly, vs) = Free(str_of_int 0, HOLogic.realT)  
  5.1715 +  | poly2term' ([(c, e) : mv_monom], vs) = monom2term ((c, e), vs)
  5.1716 +  | poly2term' ((c, e) :: ces, vs) =  
  5.1717 +    Const("op +", [HOLogic.realT, HOLogic.realT] ---> HOLogic.realT) $
  5.1718 +         poly2term (ces, vs) $ monom2term ((c, e), vs)
  5.1719 +and poly2term (xs, vs) = poly2term' (rev (sort (mv_geq LEX_) (xs)), vs);
  5.1720 +
  5.1721 +
  5.1722 +(*. converts a monom into term representation .*)
  5.1723 +(*. ignores the sign of the coefficients => use only for exp-poly functions .*)
  5.1724 +fun monom2term2((c,e):mv_monom, v:string list) =  
  5.1725 +    if c=0 then Free(str_of_int 0,HOLogic.realT)  
  5.1726 +    else
  5.1727 +	(
  5.1728 +	 if list_is_null(e) then
  5.1729 +	     ( 
  5.1730 +	      Free(str_of_int (abs(c)),HOLogic.realT)  
  5.1731 +	      )
  5.1732 +	 else
  5.1733 +	     (
  5.1734 +	      if abs(c)=1 then 
  5.1735 +		  (
  5.1736 +		   powerproduct2term(e,v)
  5.1737 +		   )
  5.1738 +	      else
  5.1739 +		  (
  5.1740 +		   Const("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $
  5.1741 +		   Free(str_of_int (abs(c)),HOLogic.realT)  $
  5.1742 +		   powerproduct2term(e,v)
  5.1743 +		   )
  5.1744 +		  )
  5.1745 +	     );
  5.1746 +
  5.1747 +(*. converts the expanded polynomial representation into the term representation .*)
  5.1748 +fun exp2term' ([]:mv_poly,vars) =  Free(str_of_int 0,HOLogic.realT)  
  5.1749 +  | exp2term' ([(c,e)],vars) =     monom2term((c,e),vars) 			     
  5.1750 +  | exp2term' ((c1,e1)::others,vars) =  
  5.1751 +    if c1<0 then 	
  5.1752 +	Const("op -",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $
  5.1753 +	exp2term'(others,vars) $
  5.1754 +	( 
  5.1755 +	 monom2term2((c1,e1),vars)
  5.1756 +	 ) 
  5.1757 +    else
  5.1758 +	Const("op +",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $
  5.1759 +	exp2term'(others,vars) $
  5.1760 +	( 
  5.1761 +	 monom2term2((c1,e1),vars)
  5.1762 +	 );
  5.1763 +	
  5.1764 +(*. sorts the powerproduct by lexicographic termorder and converts them into 
  5.1765 +    a term in polynomial representation .*)
  5.1766 +fun poly2expanded (xs,vars) = exp2term'(rev(sort (mv_geq LEX_) (xs)),vars);
  5.1767 +
  5.1768 +(*. converts a polynomial into expanded form .*)
  5.1769 +fun polynomial2expanded t =  
  5.1770 +    (let 
  5.1771 +	val vars=(((map free2str) o vars) t);
  5.1772 +    in
  5.1773 +	SOME (poly2expanded (the (term2poly t vars), vars))
  5.1774 +    end) handle _ => NONE;
  5.1775 +
  5.1776 +(*. converts a polynomial into polynomial form .*)
  5.1777 +fun expanded2polynomial t =  
  5.1778 +    (let 
  5.1779 +	val vars=(((map free2str) o vars) t);
  5.1780 +    in
  5.1781 +	SOME (poly2term (the (expanded2poly t vars), vars))
  5.1782 +    end) handle _ => NONE;
  5.1783 +
  5.1784 +
  5.1785 +(*. calculates the greatest common divisor of numerator and denominator and seperates it from each .*)
  5.1786 +fun step_cancel (t as Const ("HOL.divide",_) $ p1 $ p2) = 
  5.1787 +    let
  5.1788 +	val p1' = ref [];
  5.1789 +	val p2' = ref [];
  5.1790 +	val p3  = ref []
  5.1791 +	val vars = rev(get_vars(p1) union get_vars(p2));
  5.1792 +    in
  5.1793 +	(
  5.1794 +         p1':= sort (mv_geq LEX_) (the (term2poly p1 vars ));
  5.1795 +       	 p2':= sort (mv_geq LEX_) (the (term2poly p2 vars ));
  5.1796 +	 p3:= sort (mv_geq LEX_) (mv_gcd (!p1') (!p2'));
  5.1797 +	 if (!p3)=[(1,mv_null2(vars))] then 
  5.1798 +	     (
  5.1799 +	      Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ p1 $ p2
  5.1800 +	      )
  5.1801 +	 else
  5.1802 +	     (
  5.1803 +
  5.1804 +	      p1':=sort (mv_geq LEX_) (#1(mv_division((!p1'),(!p3),LEX_)));
  5.1805 +	      p2':=sort (mv_geq LEX_) (#1(mv_division((!p2'),(!p3),LEX_)));
  5.1806 +	      
  5.1807 +	      if #1(hd(sort (mv_geq LEX_) (!p2'))) (*mv_lc2(!p2',LEX_)*)>0 then
  5.1808 +	      (
  5.1809 +	       Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.1810 +	       $ 
  5.1811 +	       (
  5.1812 +		Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1813 +		poly2term(!p1',vars) $ 
  5.1814 +		poly2term(!p3,vars) 
  5.1815 +		) 
  5.1816 +	       $ 
  5.1817 +	       (
  5.1818 +		Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1819 +		poly2term(!p2',vars) $ 
  5.1820 +		poly2term(!p3,vars)
  5.1821 +		) 	
  5.1822 +	       )	
  5.1823 +	      else
  5.1824 +	      (
  5.1825 +	       p1':=mv_skalar_mul(!p1',~1);
  5.1826 +	       p2':=mv_skalar_mul(!p2',~1);
  5.1827 +	       p3:=mv_skalar_mul(!p3,~1);
  5.1828 +	       (
  5.1829 +		Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.1830 +		$ 
  5.1831 +		(
  5.1832 +		 Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1833 +		 poly2term(!p1',vars) $ 
  5.1834 +		 poly2term(!p3,vars) 
  5.1835 +		 ) 
  5.1836 +		$ 
  5.1837 +		(
  5.1838 +		 Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1839 +		 poly2term(!p2',vars) $ 
  5.1840 +		 poly2term(!p3,vars)
  5.1841 +		 ) 	
  5.1842 +		)	
  5.1843 +	       )	  
  5.1844 +	      )
  5.1845 +	     )
  5.1846 +    end
  5.1847 +| step_cancel _ = raise error ("RATIONALS_STEP_CANCEL_EXCEPTION: Invalid fraction"); 
  5.1848 +
  5.1849 +
  5.1850 +(*. same as step_cancel, this time for expanded forms (input+output) .*)
  5.1851 +fun step_cancel_expanded (t as Const ("HOL.divide",_) $ p1 $ p2) = 
  5.1852 +    let
  5.1853 +	val p1' = ref [];
  5.1854 +	val p2' = ref [];
  5.1855 +	val p3  = ref []
  5.1856 +	val vars = rev(get_vars(p1) union get_vars(p2));
  5.1857 +    in
  5.1858 +	(
  5.1859 +         p1':= sort (mv_geq LEX_) (the (expanded2poly p1 vars ));
  5.1860 +       	 p2':= sort (mv_geq LEX_) (the (expanded2poly p2 vars ));
  5.1861 +	 p3:= sort (mv_geq LEX_) (mv_gcd (!p1') (!p2'));
  5.1862 +	 if (!p3)=[(1,mv_null2(vars))] then 
  5.1863 +	     (
  5.1864 +	      Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ p1 $ p2
  5.1865 +	      )
  5.1866 +	 else
  5.1867 +	     (
  5.1868 +
  5.1869 +	      p1':=sort (mv_geq LEX_) (#1(mv_division((!p1'),(!p3),LEX_)));
  5.1870 +	      p2':=sort (mv_geq LEX_) (#1(mv_division((!p2'),(!p3),LEX_)));
  5.1871 +	      
  5.1872 +	      if #1(hd(sort (mv_geq LEX_) (!p2')))(* mv_lc2(!p2',LEX_)*)>0 then
  5.1873 +	      (
  5.1874 +	       Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.1875 +	       $ 
  5.1876 +	       (
  5.1877 +		Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1878 +		poly2expanded(!p1',vars) $ 
  5.1879 +		poly2expanded(!p3,vars) 
  5.1880 +		) 
  5.1881 +	       $ 
  5.1882 +	       (
  5.1883 +		Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1884 +		poly2expanded(!p2',vars) $ 
  5.1885 +		poly2expanded(!p3,vars)
  5.1886 +		) 	
  5.1887 +	       )	
  5.1888 +	      else
  5.1889 +	      (
  5.1890 +	       p1':=mv_skalar_mul(!p1',~1);
  5.1891 +	       p2':=mv_skalar_mul(!p2',~1);
  5.1892 +	       p3:=mv_skalar_mul(!p3,~1);
  5.1893 +	       (
  5.1894 +		Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.1895 +		$ 
  5.1896 +		(
  5.1897 +		 Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1898 +		 poly2expanded(!p1',vars) $ 
  5.1899 +		 poly2expanded(!p3,vars) 
  5.1900 +		 ) 
  5.1901 +		$ 
  5.1902 +		(
  5.1903 +		 Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.1904 +		 poly2expanded(!p2',vars) $ 
  5.1905 +		 poly2expanded(!p3,vars)
  5.1906 +		 ) 	
  5.1907 +		)	
  5.1908 +	       )	  
  5.1909 +	      )
  5.1910 +	     )
  5.1911 +    end
  5.1912 +| step_cancel_expanded _ = raise error ("RATIONALS_STEP_CANCEL_EXCEPTION: Invalid fraction"); 
  5.1913 +
  5.1914 +(*. calculates the greatest common divisor of numerator and denominator and divides each through it .*)
  5.1915 +fun direct_cancel (t as Const ("HOL.divide",_) $ p1 $ p2) = 
  5.1916 +    let
  5.1917 +	val p1' = ref [];
  5.1918 +	val p2' = ref [];
  5.1919 +	val p3  = ref []
  5.1920 +	val vars = rev(get_vars(p1) union get_vars(p2));
  5.1921 +    in
  5.1922 +	(
  5.1923 +	 p1':=sort (mv_geq LEX_) (mv_shorten((the (term2poly p1 vars )),LEX_));
  5.1924 +	 p2':=sort (mv_geq LEX_) (mv_shorten((the (term2poly p2 vars )),LEX_));	 
  5.1925 +	 p3 :=sort (mv_geq LEX_) (mv_gcd (!p1') (!p2'));
  5.1926 +
  5.1927 +	 if (!p3)=[(1,mv_null2(vars))] then 
  5.1928 +	     (
  5.1929 +	      (Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ p1 $ p2,[])
  5.1930 +	      )
  5.1931 +	 else
  5.1932 +	     (
  5.1933 +	      p1':=sort (mv_geq LEX_) (#1(mv_division((!p1'),(!p3),LEX_)));
  5.1934 +	      p2':=sort (mv_geq LEX_) (#1(mv_division((!p2'),(!p3),LEX_)));
  5.1935 +	      if #1(hd(sort (mv_geq LEX_) (!p2'))) (*mv_lc2(!p2',LEX_)*)>0 then	      
  5.1936 +	      (
  5.1937 +	       (
  5.1938 +		Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.1939 +		$ 
  5.1940 +		(
  5.1941 +		 poly2term((!p1'),vars)
  5.1942 +		 ) 
  5.1943 +		$ 
  5.1944 +		( 
  5.1945 +		 poly2term((!p2'),vars)
  5.1946 +		 ) 	
  5.1947 +		)
  5.1948 +	       ,
  5.1949 +	       if mv_grad(!p3)>0 then 
  5.1950 +		   [
  5.1951 +		    (
  5.1952 +		     Const ("Not",[bool]--->bool) $
  5.1953 +		     (
  5.1954 +		      Const("op =",[HOLogic.realT,HOLogic.realT]--->bool) $
  5.1955 +		      poly2term((!p3),vars) $
  5.1956 +		      Free("0",HOLogic.realT)
  5.1957 +		      )
  5.1958 +		     )
  5.1959 +		    ]
  5.1960 +	       else
  5.1961 +		   []
  5.1962 +		   )
  5.1963 +	      else
  5.1964 +		  (
  5.1965 +		   p1':=mv_skalar_mul(!p1',~1);
  5.1966 +		   p2':=mv_skalar_mul(!p2',~1);
  5.1967 +		   if length(!p3)> 2*(count_neg(!p3)) then () else p3 :=mv_skalar_mul(!p3,~1); 
  5.1968 +		       (
  5.1969 +			Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.1970 +			$ 
  5.1971 +			(
  5.1972 +			 poly2term((!p1'),vars)
  5.1973 +			 ) 
  5.1974 +			$ 
  5.1975 +			( 
  5.1976 +			 poly2term((!p2'),vars)
  5.1977 +			 ) 	
  5.1978 +			,
  5.1979 +			if mv_grad(!p3)>0 then 
  5.1980 +			    [
  5.1981 +			     (
  5.1982 +			      Const ("Not",[bool]--->bool) $
  5.1983 +			      (
  5.1984 +			       Const("op =",[HOLogic.realT,HOLogic.realT]--->bool) $
  5.1985 +			       poly2term((!p3),vars) $
  5.1986 +			       Free("0",HOLogic.realT)
  5.1987 +			       )
  5.1988 +			      )
  5.1989 +			     ]
  5.1990 +			else
  5.1991 +			    []
  5.1992 +			    )
  5.1993 +		       )
  5.1994 +		  )
  5.1995 +	     )
  5.1996 +    end
  5.1997 +  | direct_cancel _ = raise error ("RATIONALS_DIRECT_CANCEL_EXCEPTION: Invalid fraction"); 
  5.1998 +
  5.1999 +(*. same es direct_cancel, this time for expanded forms (input+output).*) 
  5.2000 +fun direct_cancel_expanded (t as Const ("HOL.divide",_) $ p1 $ p2) =  
  5.2001 +    let
  5.2002 +	val p1' = ref [];
  5.2003 +	val p2' = ref [];
  5.2004 +	val p3  = ref []
  5.2005 +	val vars = rev(get_vars(p1) union get_vars(p2));
  5.2006 +    in
  5.2007 +	(
  5.2008 +	 p1':=sort (mv_geq LEX_) (mv_shorten((the (expanded2poly p1 vars )),LEX_));
  5.2009 +	 p2':=sort (mv_geq LEX_) (mv_shorten((the (expanded2poly p2 vars )),LEX_));	 
  5.2010 +	 p3 :=sort (mv_geq LEX_) (mv_gcd (!p1') (!p2'));
  5.2011 +
  5.2012 +	 if (!p3)=[(1,mv_null2(vars))] then 
  5.2013 +	     (
  5.2014 +	      (Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ p1 $ p2,[])
  5.2015 +	      )
  5.2016 +	 else
  5.2017 +	     (
  5.2018 +	      p1':=sort (mv_geq LEX_) (#1(mv_division((!p1'),(!p3),LEX_)));
  5.2019 +	      p2':=sort (mv_geq LEX_) (#1(mv_division((!p2'),(!p3),LEX_)));
  5.2020 +	      if #1(hd(sort (mv_geq LEX_) (!p2'))) (*mv_lc2(!p2',LEX_)*)>0 then	      
  5.2021 +	      (
  5.2022 +	       (
  5.2023 +		Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2024 +		$ 
  5.2025 +		(
  5.2026 +		 poly2expanded((!p1'),vars)
  5.2027 +		 ) 
  5.2028 +		$ 
  5.2029 +		( 
  5.2030 +		 poly2expanded((!p2'),vars)
  5.2031 +		 ) 	
  5.2032 +		)
  5.2033 +	       ,
  5.2034 +	       if mv_grad(!p3)>0 then 
  5.2035 +		   [
  5.2036 +		    (
  5.2037 +		     Const ("Not",[bool]--->bool) $
  5.2038 +		     (
  5.2039 +		      Const("op =",[HOLogic.realT,HOLogic.realT]--->bool) $
  5.2040 +		      poly2expanded((!p3),vars) $
  5.2041 +		      Free("0",HOLogic.realT)
  5.2042 +		      )
  5.2043 +		     )
  5.2044 +		    ]
  5.2045 +	       else
  5.2046 +		   []
  5.2047 +		   )
  5.2048 +	      else
  5.2049 +		  (
  5.2050 +		   p1':=mv_skalar_mul(!p1',~1);
  5.2051 +		   p2':=mv_skalar_mul(!p2',~1);
  5.2052 +		   if length(!p3)> 2*(count_neg(!p3)) then () else p3 :=mv_skalar_mul(!p3,~1); 
  5.2053 +		       (
  5.2054 +			Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2055 +			$ 
  5.2056 +			(
  5.2057 +			 poly2expanded((!p1'),vars)
  5.2058 +			 ) 
  5.2059 +			$ 
  5.2060 +			( 
  5.2061 +			 poly2expanded((!p2'),vars)
  5.2062 +			 ) 	
  5.2063 +			,
  5.2064 +			if mv_grad(!p3)>0 then 
  5.2065 +			    [
  5.2066 +			     (
  5.2067 +			      Const ("Not",[bool]--->bool) $
  5.2068 +			      (
  5.2069 +			       Const("op =",[HOLogic.realT,HOLogic.realT]--->bool) $
  5.2070 +			       poly2expanded((!p3),vars) $
  5.2071 +			       Free("0",HOLogic.realT)
  5.2072 +			       )
  5.2073 +			      )
  5.2074 +			     ]
  5.2075 +			else
  5.2076 +			    []
  5.2077 +			    )
  5.2078 +		       )
  5.2079 +		  )
  5.2080 +	     )
  5.2081 +    end
  5.2082 +  | direct_cancel_expanded _ = raise error ("RATIONALS_DIRECT_CANCEL_EXCEPTION: Invalid fraction"); 
  5.2083 +
  5.2084 +
  5.2085 +(*. adds two fractions .*)
  5.2086 +fun add_fract ((Const("HOL.divide",_) $ t11 $ t12),(Const("HOL.divide",_) $ t21 $ t22)) =
  5.2087 +    let
  5.2088 +	val vars=get_vars(t11) union get_vars(t12) union get_vars(t21) union get_vars(t22);
  5.2089 +	val t11'=ref (the(term2poly t11 vars));
  5.2090 +val _= writeln"### add_fract: done t11"
  5.2091 +	val t12'=ref (the(term2poly t12 vars));
  5.2092 +val _= writeln"### add_fract: done t12"
  5.2093 +	val t21'=ref (the(term2poly t21 vars));
  5.2094 +val _= writeln"### add_fract: done t21"
  5.2095 +	val t22'=ref (the(term2poly t22 vars));
  5.2096 +val _= writeln"### add_fract: done t22"
  5.2097 +	val den=ref [];
  5.2098 +	val nom=ref [];
  5.2099 +	val m1=ref [];
  5.2100 +	val m2=ref [];
  5.2101 +    in
  5.2102 +	
  5.2103 +	(
  5.2104 +	 den :=sort (mv_geq LEX_) (mv_lcm (!t12') (!t22'));
  5.2105 +writeln"### add_fract: done sort mv_lcm";
  5.2106 +	 m1  :=sort (mv_geq LEX_) (#1(mv_division(!den,!t12',LEX_)));
  5.2107 +writeln"### add_fract: done sort mv_division t12";
  5.2108 +	 m2  :=sort (mv_geq LEX_) (#1(mv_division(!den,!t22',LEX_)));
  5.2109 +writeln"### add_fract: done sort mv_division t22";
  5.2110 +	 nom :=sort (mv_geq LEX_) 
  5.2111 +		    (mv_shorten(mv_add(mv_mul(!t11',!m1,LEX_),
  5.2112 +				       mv_mul(!t21',!m2,LEX_),
  5.2113 +				       LEX_),
  5.2114 +				LEX_));
  5.2115 +writeln"### add_fract: done sort mv_add";
  5.2116 +	 (
  5.2117 +	  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2118 +	  $ 
  5.2119 +	  (
  5.2120 +	   poly2term((!nom),vars)
  5.2121 +	   ) 
  5.2122 +	  $ 
  5.2123 +	  ( 
  5.2124 +	   poly2term((!den),vars)
  5.2125 +	   )	      
  5.2126 +	  )
  5.2127 +	 )	     
  5.2128 +    end 
  5.2129 +  | add_fract (_,_) = raise error ("RATIONALS_ADD_FRACTION_EXCEPTION: Invalid add_fraction call");
  5.2130 +
  5.2131 +(*. adds two expanded fractions .*)
  5.2132 +fun add_fract_exp ((Const("HOL.divide",_) $ t11 $ t12),(Const("HOL.divide",_) $ t21 $ t22)) =
  5.2133 +    let
  5.2134 +	val vars=get_vars(t11) union get_vars(t12) union get_vars(t21) union get_vars(t22);
  5.2135 +	val t11'=ref (the(expanded2poly t11 vars));
  5.2136 +	val t12'=ref (the(expanded2poly t12 vars));
  5.2137 +	val t21'=ref (the(expanded2poly t21 vars));
  5.2138 +	val t22'=ref (the(expanded2poly t22 vars));
  5.2139 +	val den=ref [];
  5.2140 +	val nom=ref [];
  5.2141 +	val m1=ref [];
  5.2142 +	val m2=ref [];
  5.2143 +    in
  5.2144 +	
  5.2145 +	(
  5.2146 +	 den :=sort (mv_geq LEX_) (mv_lcm (!t12') (!t22'));
  5.2147 +	 m1  :=sort (mv_geq LEX_) (#1(mv_division(!den,!t12',LEX_)));
  5.2148 +	 m2  :=sort (mv_geq LEX_) (#1(mv_division(!den,!t22',LEX_)));
  5.2149 +	 nom :=sort (mv_geq LEX_) (mv_shorten(mv_add(mv_mul(!t11',!m1,LEX_),mv_mul(!t21',!m2,LEX_),LEX_),LEX_));
  5.2150 +	 (
  5.2151 +	  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2152 +	  $ 
  5.2153 +	  (
  5.2154 +	   poly2expanded((!nom),vars)
  5.2155 +	   ) 
  5.2156 +	  $ 
  5.2157 +	  ( 
  5.2158 +	   poly2expanded((!den),vars)
  5.2159 +	   )	      
  5.2160 +	  )
  5.2161 +	 )	     
  5.2162 +    end 
  5.2163 +  | add_fract_exp (_,_) = raise error ("RATIONALS_ADD_FRACTION_EXP_EXCEPTION: Invalid add_fraction call");
  5.2164 +
  5.2165 +(*. adds a list of terms .*)
  5.2166 +fun add_list_of_fractions []= (Free("0",HOLogic.realT),[])
  5.2167 +  | add_list_of_fractions [x]= direct_cancel x
  5.2168 +  | add_list_of_fractions (x::y::xs) = 
  5.2169 +    let
  5.2170 +	val (t1a,rest1)=direct_cancel(x);
  5.2171 +val _= writeln"### add_list_of_fractions xs: has done direct_cancel(x)";
  5.2172 +	val (t2a,rest2)=direct_cancel(y);
  5.2173 +val _= writeln"### add_list_of_fractions xs: has done direct_cancel(y)";
  5.2174 +	val (t3a,rest3)=(add_list_of_fractions (add_fract(t1a,t2a)::xs));
  5.2175 +val _= writeln"### add_list_of_fractions xs: has done add_list_of_fraction xs";
  5.2176 +	val (t4a,rest4)=direct_cancel(t3a);
  5.2177 +val _= writeln"### add_list_of_fractions xs: has done direct_cancel(t3a)";
  5.2178 +	val rest=rest1 union rest2 union rest3 union rest4;
  5.2179 +    in
  5.2180 +	(writeln"### add_list_of_fractions in";
  5.2181 +	 (
  5.2182 +	 (t4a,rest) 
  5.2183 +	 )
  5.2184 +	 )
  5.2185 +    end;
  5.2186 +
  5.2187 +(*. adds a list of expanded terms .*)
  5.2188 +fun add_list_of_fractions_exp []= (Free("0",HOLogic.realT),[])
  5.2189 +  | add_list_of_fractions_exp [x]= direct_cancel_expanded x
  5.2190 +  | add_list_of_fractions_exp (x::y::xs) = 
  5.2191 +    let
  5.2192 +	val (t1a,rest1)=direct_cancel_expanded(x);
  5.2193 +	val (t2a,rest2)=direct_cancel_expanded(y);
  5.2194 +	val (t3a,rest3)=(add_list_of_fractions_exp (add_fract_exp(t1a,t2a)::xs));
  5.2195 +	val (t4a,rest4)=direct_cancel_expanded(t3a);
  5.2196 +	val rest=rest1 union rest2 union rest3 union rest4;
  5.2197 +    in
  5.2198 +	(
  5.2199 +	 (t4a,rest) 
  5.2200 +	 )
  5.2201 +    end;
  5.2202 +
  5.2203 +(*. calculates the lcm of a list of mv_poly .*)
  5.2204 +fun calc_lcm ([x],var)= (x,var) 
  5.2205 +  | calc_lcm ((x::xs),var) = (mv_lcm x (#1(calc_lcm (xs,var))),var);
  5.2206 +
  5.2207 +(*. converts a list of terms to a list of mv_poly .*)
  5.2208 +fun t2d([],_)=[] 
  5.2209 +  | t2d((t as (Const("HOL.divide",_) $ p1 $ p2))::xs,vars)= (the(term2poly p2 vars)) :: t2d(xs,vars); 
  5.2210 +
  5.2211 +(*. same as t2d, this time for expanded forms .*)
  5.2212 +fun t2d_exp([],_)=[]  
  5.2213 +  | t2d_exp((t as (Const("HOL.divide",_) $ p1 $ p2))::xs,vars)= (the(expanded2poly p2 vars)) :: t2d_exp(xs,vars);
  5.2214 +
  5.2215 +(*. converts a list of fract terms to a list of their denominators .*)
  5.2216 +fun termlist2denominators [] = ([],[])
  5.2217 +  | termlist2denominators xs = 
  5.2218 +    let	
  5.2219 +	val xxs=ref xs;
  5.2220 +	val var=ref [];
  5.2221 +    in
  5.2222 +	var:=[];
  5.2223 +	while length(!xxs)>0 do
  5.2224 +	    (
  5.2225 +	     let 
  5.2226 +		 val (t as Const ("HOL.divide",_) $ p1x $ p2x)=hd(!xxs);
  5.2227 +	     in
  5.2228 +		 (
  5.2229 +		  xxs:=tl(!xxs);
  5.2230 +		  var:=((get_vars(p2x)) union (get_vars(p1x)) union (!var))
  5.2231 +		  )
  5.2232 +	     end
  5.2233 +	     );
  5.2234 +	    (t2d(xs,!var),!var)
  5.2235 +    end;
  5.2236 +
  5.2237 +(*. calculates the lcm of a list of mv_poly .*)
  5.2238 +fun calc_lcm ([x],var)= (x,var) 
  5.2239 +  | calc_lcm ((x::xs),var) = (mv_lcm x (#1(calc_lcm (xs,var))),var);
  5.2240 +
  5.2241 +(*. converts a list of terms to a list of mv_poly .*)
  5.2242 +fun t2d([],_)=[] 
  5.2243 +  | t2d((t as (Const("HOL.divide",_) $ p1 $ p2))::xs,vars)= (the(term2poly p2 vars)) :: t2d(xs,vars); 
  5.2244 +
  5.2245 +(*. same as t2d, this time for expanded forms .*)
  5.2246 +fun t2d_exp([],_)=[]  
  5.2247 +  | t2d_exp((t as (Const("HOL.divide",_) $ p1 $ p2))::xs,vars)= (the(expanded2poly p2 vars)) :: t2d_exp(xs,vars);
  5.2248 +
  5.2249 +(*. converts a list of fract terms to a list of their denominators .*)
  5.2250 +fun termlist2denominators [] = ([],[])
  5.2251 +  | termlist2denominators xs = 
  5.2252 +    let	
  5.2253 +	val xxs=ref xs;
  5.2254 +	val var=ref [];
  5.2255 +    in
  5.2256 +	var:=[];
  5.2257 +	while length(!xxs)>0 do
  5.2258 +	    (
  5.2259 +	     let 
  5.2260 +		 val (t as Const ("HOL.divide",_) $ p1x $ p2x)=hd(!xxs);
  5.2261 +	     in
  5.2262 +		 (
  5.2263 +		  xxs:=tl(!xxs);
  5.2264 +		  var:=((get_vars(p2x)) union (get_vars(p1x)) union (!var))
  5.2265 +		  )
  5.2266 +	     end
  5.2267 +	     );
  5.2268 +	    (t2d(xs,!var),!var)
  5.2269 +    end;
  5.2270 +
  5.2271 +(*. same as termlist2denminators, this time for expanded forms .*)
  5.2272 +fun termlist2denominators_exp [] = ([],[])
  5.2273 +  | termlist2denominators_exp xs = 
  5.2274 +    let	
  5.2275 +	val xxs=ref xs;
  5.2276 +	val var=ref [];
  5.2277 +    in
  5.2278 +	var:=[];
  5.2279 +	while length(!xxs)>0 do
  5.2280 +	    (
  5.2281 +	     let 
  5.2282 +		 val (t as Const ("HOL.divide",_) $ p1x $ p2x)=hd(!xxs);
  5.2283 +	     in
  5.2284 +		 (
  5.2285 +		  xxs:=tl(!xxs);
  5.2286 +		  var:=((get_vars(p2x)) union (get_vars(p1x)) union (!var))
  5.2287 +		  )
  5.2288 +	     end
  5.2289 +	     );
  5.2290 +	    (t2d_exp(xs,!var),!var)
  5.2291 +    end;
  5.2292 +
  5.2293 +(*. reduces all fractions to the least common denominator .*)
  5.2294 +fun com_den(x::xs,denom,den,var)=
  5.2295 +    let 
  5.2296 +	val (t as Const ("HOL.divide",_) $ p1' $ p2')=x;
  5.2297 +	val p2= sort (mv_geq LEX_) (the(term2poly p2' var));
  5.2298 +	val p3= #1(mv_division(denom,p2,LEX_));
  5.2299 +	val p1var=get_vars(p1');
  5.2300 +    in     
  5.2301 +	if length(xs)>0 then 
  5.2302 +	    if p3=[(1,mv_null2(var))] then
  5.2303 +		(
  5.2304 +		 Const ("op +",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT)
  5.2305 +		 $ 
  5.2306 +		 (
  5.2307 +		  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2308 +		  $ 
  5.2309 +		  poly2term(the (term2poly p1' p1var),p1var)
  5.2310 +		  $ 
  5.2311 +		  den	
  5.2312 +		  )    
  5.2313 +		 $ 
  5.2314 +		 #1(com_den(xs,denom,den,var))
  5.2315 +		,
  5.2316 +		[]
  5.2317 +		)
  5.2318 +	    else
  5.2319 +		(
  5.2320 +		 Const ("op +",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2321 +		 $ 
  5.2322 +		 (
  5.2323 +		  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2324 +		  $ 
  5.2325 +		  (
  5.2326 +		   Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2327 +		   poly2term(the (term2poly p1' p1var),p1var) $ 
  5.2328 +		   poly2term(p3,var)
  5.2329 +		   ) 
  5.2330 +		  $ 
  5.2331 +		  (
  5.2332 +		   den
  5.2333 +		   ) 	
  5.2334 +		  )
  5.2335 +		 $ 
  5.2336 +		 #1(com_den(xs,denom,den,var))
  5.2337 +		,
  5.2338 +		[]
  5.2339 +		)
  5.2340 +	else
  5.2341 +	    if p3=[(1,mv_null2(var))] then
  5.2342 +		(
  5.2343 +		 (
  5.2344 +		  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2345 +		  $ 
  5.2346 +		  poly2term(the (term2poly p1' p1var),p1var)
  5.2347 +		  $ 
  5.2348 +		  den	
  5.2349 +		  )
  5.2350 +		 ,
  5.2351 +		 []
  5.2352 +		 )
  5.2353 +	     else
  5.2354 +		 (
  5.2355 +		  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2356 +		  $ 
  5.2357 +		  (
  5.2358 +		   Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2359 +		   poly2term(the (term2poly p1' p1var),p1var) $ 
  5.2360 +		   poly2term(p3,var)
  5.2361 +		   ) 
  5.2362 +		  $ 
  5.2363 +		  den 	
  5.2364 +		  ,
  5.2365 +		  []
  5.2366 +		  )
  5.2367 +    end;
  5.2368 +
  5.2369 +(*. same as com_den, this time for expanded forms .*)
  5.2370 +fun com_den_exp(x::xs,denom,den,var)=
  5.2371 +    let 
  5.2372 +	val (t as Const ("HOL.divide",_) $ p1' $ p2')=x;
  5.2373 +	val p2= sort (mv_geq LEX_) (the(expanded2poly p2' var));
  5.2374 +	val p3= #1(mv_division(denom,p2,LEX_));
  5.2375 +	val p1var=get_vars(p1');
  5.2376 +    in     
  5.2377 +	if length(xs)>0 then 
  5.2378 +	    if p3=[(1,mv_null2(var))] then
  5.2379 +		(
  5.2380 +		 Const ("op +",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT)
  5.2381 +		 $ 
  5.2382 +		 (
  5.2383 +		  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2384 +		  $ 
  5.2385 +		  poly2expanded(the(expanded2poly p1' p1var),p1var)
  5.2386 +		  $ 
  5.2387 +		  den	
  5.2388 +		  )    
  5.2389 +		 $ 
  5.2390 +		 #1(com_den_exp(xs,denom,den,var))
  5.2391 +		,
  5.2392 +		[]
  5.2393 +		)
  5.2394 +	    else
  5.2395 +		(
  5.2396 +		 Const ("op +",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2397 +		 $ 
  5.2398 +		 (
  5.2399 +		  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2400 +		  $ 
  5.2401 +		  (
  5.2402 +		   Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2403 +		   poly2expanded(the(expanded2poly p1' p1var),p1var) $ 
  5.2404 +		   poly2expanded(p3,var)
  5.2405 +		   ) 
  5.2406 +		  $ 
  5.2407 +		  (
  5.2408 +		   den
  5.2409 +		   ) 	
  5.2410 +		  )
  5.2411 +		 $ 
  5.2412 +		 #1(com_den_exp(xs,denom,den,var))
  5.2413 +		,
  5.2414 +		[]
  5.2415 +		)
  5.2416 +	else
  5.2417 +	    if p3=[(1,mv_null2(var))] then
  5.2418 +		(
  5.2419 +		 (
  5.2420 +		  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2421 +		  $ 
  5.2422 +		  poly2expanded(the(expanded2poly p1' p1var),p1var)
  5.2423 +		  $ 
  5.2424 +		  den	
  5.2425 +		  )
  5.2426 +		 ,
  5.2427 +		 []
  5.2428 +		 )
  5.2429 +	     else
  5.2430 +		 (
  5.2431 +		  Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) 
  5.2432 +		  $ 
  5.2433 +		  (
  5.2434 +		   Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2435 +		   poly2expanded(the(expanded2poly p1' p1var),p1var) $ 
  5.2436 +		   poly2expanded(p3,var)
  5.2437 +		   ) 
  5.2438 +		  $ 
  5.2439 +		  den 	
  5.2440 +		  ,
  5.2441 +		  []
  5.2442 +		  )
  5.2443 +    end;
  5.2444 +
  5.2445 +(* wird aktuell nicht mehr gebraucht, bei rückänderung schon 
  5.2446 +-------------------------------------------------------------
  5.2447 +(* WN0210???SK brauch ma des überhaupt *)
  5.2448 +fun com_den2(x::xs,denom,den,var)=
  5.2449 +    let 
  5.2450 +	val (t as Const ("HOL.divide",_) $ p1' $ p2')=x;
  5.2451 +	val p2= sort (mv_geq LEX_) (the(term2poly p2' var));
  5.2452 +	val p3= #1(mv_division(denom,p2,LEX_));
  5.2453 +	val p1var=get_vars(p1');
  5.2454 +    in     
  5.2455 +	if length(xs)>0 then 
  5.2456 +	    if p3=[(1,mv_null2(var))] then
  5.2457 +		(
  5.2458 +		 Const ("op +",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2459 +		 poly2term(the(term2poly p1' p1var),p1var) $ 
  5.2460 +		 com_den2(xs,denom,den,var)
  5.2461 +		)
  5.2462 +	    else
  5.2463 +		(
  5.2464 +		 Const ("op +",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2465 +		 (
  5.2466 +		   let 
  5.2467 +		       val p3'=poly2term(p3,var);
  5.2468 +		       val vars= (((map free2str) o vars) p1') union (((map free2str) o vars) p3');
  5.2469 +		   in
  5.2470 +		       poly2term(sort (mv_geq LEX_) (mv_mul(the(term2poly p1' vars) ,the(term2poly p3' vars),LEX_)),vars)
  5.2471 +		   end
  5.2472 +		  ) $ 
  5.2473 +		 com_den2(xs,denom,den,var)
  5.2474 +		)
  5.2475 +	else
  5.2476 +	    if p3=[(1,mv_null2(var))] then
  5.2477 +		(
  5.2478 +		 poly2term(the(term2poly p1' p1var),p1var)
  5.2479 +		 )
  5.2480 +	     else
  5.2481 +		 (
  5.2482 +		   let 
  5.2483 +		       val p3'=poly2term(p3,var);
  5.2484 +		       val vars= (((map free2str) o vars) p1') union (((map free2str) o vars) p3');
  5.2485 +		   in
  5.2486 +		       poly2term(sort (mv_geq LEX_) (mv_mul(the(term2poly p1' vars) ,the(term2poly p3' vars),LEX_)),vars)
  5.2487 +		   end
  5.2488 +		  )
  5.2489 +    end;
  5.2490 +
  5.2491 +(* WN0210???SK brauch ma des überhaupt *)
  5.2492 +fun com_den_exp2(x::xs,denom,den,var)=
  5.2493 +    let 
  5.2494 +	val (t as Const ("HOL.divide",_) $ p1' $ p2')=x;
  5.2495 +	val p2= sort (mv_geq LEX_) (the(expanded2poly p2' var));
  5.2496 +	val p3= #1(mv_division(denom,p2,LEX_));
  5.2497 +	val p1var=get_vars p1';
  5.2498 +    in     
  5.2499 +	if length(xs)>0 then 
  5.2500 +	    if p3=[(1,mv_null2(var))] then
  5.2501 +		(
  5.2502 +		 Const ("op +",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2503 +		 poly2expanded(the (expanded2poly p1' p1var),p1var) $ 
  5.2504 +		 com_den_exp2(xs,denom,den,var)
  5.2505 +		)
  5.2506 +	    else
  5.2507 +		(
  5.2508 +		 Const ("op +",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2509 +		 (
  5.2510 +		   let 
  5.2511 +		       val p3'=poly2expanded(p3,var);
  5.2512 +		       val vars= (((map free2str) o vars) p1') union (((map free2str) o vars) p3');
  5.2513 +		   in
  5.2514 +		       poly2expanded(sort (mv_geq LEX_) (mv_mul(the(expanded2poly p1' vars) ,the(expanded2poly p3' vars),LEX_)),vars)
  5.2515 +		   end
  5.2516 +		  ) $ 
  5.2517 +		 com_den_exp2(xs,denom,den,var)
  5.2518 +		)
  5.2519 +	else
  5.2520 +	    if p3=[(1,mv_null2(var))] then
  5.2521 +		(
  5.2522 +		 poly2expanded(the (expanded2poly p1' p1var),p1var)
  5.2523 +		 )
  5.2524 +	     else
  5.2525 +		 (
  5.2526 +		   let 
  5.2527 +		       val p3'=poly2expanded(p3,var);
  5.2528 +		       val vars= (((map free2str) o vars) p1') union (((map free2str) o vars) p3');
  5.2529 +		   in
  5.2530 +		       poly2expanded(sort (mv_geq LEX_) (mv_mul(the(expanded2poly p1' vars) ,the(expanded2poly p3' vars),LEX_)),vars)
  5.2531 +		   end
  5.2532 +		  )
  5.2533 +    end;
  5.2534 +---------------------------------------------------------*)
  5.2535 +
  5.2536 +
  5.2537 +(*. searches for an element y of a list ys, which has an gcd not 1 with x .*) 
  5.2538 +fun exists_gcd (x,[]) = false 
  5.2539 +  | exists_gcd (x,y::ys) = if mv_gcd x y = [(1,mv_null2(#2(hd(x))))] then  exists_gcd (x,ys)
  5.2540 +			   else true;
  5.2541 +
  5.2542 +(*. divides each element of the list xs with y .*)
  5.2543 +fun list_div ([],y) = [] 
  5.2544 +  | list_div (x::xs,y) = 
  5.2545 +    let
  5.2546 +	val (d,r)=mv_division(x,y,LEX_);
  5.2547 +    in
  5.2548 +	if r=[] then 
  5.2549 +	    d::list_div(xs,y)
  5.2550 +	else x::list_div(xs,y)
  5.2551 +    end;
  5.2552 +    
  5.2553 +(*. checks if x is in the list ys .*)
  5.2554 +fun in_list (x,[]) = false 
  5.2555 +  | in_list (x,y::ys) = if x=y then true
  5.2556 +			else in_list(x,ys);
  5.2557 +
  5.2558 +(*. deletes all equal elements of the list xs .*)
  5.2559 +fun kill_equal [] = [] 
  5.2560 +  | kill_equal (x::xs) = if in_list(x,xs) orelse x=[(1,mv_null2(#2(hd(x))))] then kill_equal(xs)
  5.2561 +			 else x::kill_equal(xs);
  5.2562 +
  5.2563 +(*. searches for new factors .*)
  5.2564 +fun new_factors [] = []
  5.2565 +  | new_factors (list:mv_poly list):mv_poly list = 
  5.2566 +    let
  5.2567 +	val l = kill_equal list;
  5.2568 +	val len = length(l);
  5.2569 +    in
  5.2570 +	if len>=2 then
  5.2571 +	    (
  5.2572 +	     let
  5.2573 +		 val x::y::xs=l;
  5.2574 +		 val gcd=mv_gcd x y;
  5.2575 +	     in
  5.2576 +		 if gcd=[(1,mv_null2(#2(hd(x))))] then 
  5.2577 +		     ( 
  5.2578 +		      if exists_gcd(x,xs) then new_factors (y::xs @ [x])
  5.2579 +		      else x::new_factors(y::xs)
  5.2580 +	             )
  5.2581 +		 else gcd::new_factors(kill_equal(list_div(x::y::xs,gcd)))
  5.2582 +	     end
  5.2583 +	     )
  5.2584 +	else
  5.2585 +	    if len=1 then [hd(l)]
  5.2586 +	    else []
  5.2587 +    end;
  5.2588 +
  5.2589 +(*. gets the factors of a list .*)
  5.2590 +fun get_factors x = new_factors x; 
  5.2591 +
  5.2592 +(*. multiplies the elements of the list .*)
  5.2593 +fun multi_list [] = []
  5.2594 +  | multi_list (x::xs) = if xs=[] then x
  5.2595 +			 else mv_mul(x,multi_list xs,LEX_);
  5.2596 +
  5.2597 +(*. makes a term out of the elements of the list (polynomial representation) .*)
  5.2598 +fun make_term ([],vars) = Free(str_of_int 0,HOLogic.realT) 
  5.2599 +  | make_term ((x::xs),vars) = if length(xs)=0 then poly2term(sort (mv_geq LEX_) (x),vars)
  5.2600 +			       else
  5.2601 +				   (
  5.2602 +				    Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2603 +				    poly2term(sort (mv_geq LEX_) (x),vars) $ 
  5.2604 +				    make_term(xs,vars)
  5.2605 +				    );
  5.2606 +
  5.2607 +(*. factorizes the denominator (polynomial representation) .*)				
  5.2608 +fun factorize_den (l,den,vars) = 
  5.2609 +    let
  5.2610 +	val factor_list=kill_equal( (get_factors l));
  5.2611 +	val mlist=multi_list(factor_list);
  5.2612 +	val (last,rest)=mv_division(den,multi_list(factor_list),LEX_);
  5.2613 +    in
  5.2614 +	if rest=[] then
  5.2615 +	    (
  5.2616 +	     if last=[(1,mv_null2(vars))] then make_term(factor_list,vars)
  5.2617 +	     else make_term(last::factor_list,vars)
  5.2618 +	     )
  5.2619 +	else raise error ("RATIONALS_FACTORIZE_DEN_EXCEPTION: Invalid factor by division")
  5.2620 +    end; 
  5.2621 +
  5.2622 +(*. makes a term out of the elements of the list (expanded polynomial representation) .*)
  5.2623 +fun make_exp ([],vars) = Free(str_of_int 0,HOLogic.realT) 
  5.2624 +  | make_exp ((x::xs),vars) = if length(xs)=0 then poly2expanded(sort (mv_geq LEX_) (x),vars)
  5.2625 +			       else
  5.2626 +				   (
  5.2627 +				    Const ("op *",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2628 +				    poly2expanded(sort (mv_geq LEX_) (x),vars) $ 
  5.2629 +				    make_exp(xs,vars)
  5.2630 +				    );
  5.2631 +
  5.2632 +(*. factorizes the denominator (expanded polynomial representation) .*)	
  5.2633 +fun factorize_den_exp (l,den,vars) = 
  5.2634 +    let
  5.2635 +	val factor_list=kill_equal( (get_factors l));
  5.2636 +	val mlist=multi_list(factor_list);
  5.2637 +	val (last,rest)=mv_division(den,multi_list(factor_list),LEX_);
  5.2638 +    in
  5.2639 +	if rest=[] then
  5.2640 +	    (
  5.2641 +	     if last=[(1,mv_null2(vars))] then make_exp(factor_list,vars)
  5.2642 +	     else make_exp(last::factor_list,vars)
  5.2643 +	     )
  5.2644 +	else raise error ("RATIONALS_FACTORIZE_DEN_EXP_EXCEPTION: Invalid factor by division")
  5.2645 +    end; 
  5.2646 +
  5.2647 +(*. calculates the common denominator of all elements of the list and multiplies .*)
  5.2648 +(*. the nominators and denominators with the correct factor .*)
  5.2649 +(*. (polynomial representation) .*)
  5.2650 +fun step_add_list_of_fractions []=(Free("0",HOLogic.realT),[]:term list)
  5.2651 +  | step_add_list_of_fractions [x]= raise error ("RATIONALS_STEP_ADD_LIST_OF_FRACTIONS_EXCEPTION: Nothing to add")
  5.2652 +  | step_add_list_of_fractions (xs) = 
  5.2653 +    let
  5.2654 +        val den_list=termlist2denominators (xs); (* list of denominators *)
  5.2655 +	val (denom,var)=calc_lcm(den_list);      (* common denominator *)
  5.2656 +	val den=factorize_den(#1(den_list),denom,var); (* faktorisierter Nenner !!! *)
  5.2657 +    in
  5.2658 +	com_den(xs,denom,den,var)
  5.2659 +    end;
  5.2660 +
  5.2661 +(*. calculates the common denominator of all elements of the list and multiplies .*)
  5.2662 +(*. the nominators and denominators with the correct factor .*)
  5.2663 +(*. (expanded polynomial representation) .*)
  5.2664 +fun step_add_list_of_fractions_exp []  = (Free("0",HOLogic.realT),[]:term list)
  5.2665 +  | step_add_list_of_fractions_exp [x] = raise error ("RATIONALS_STEP_ADD_LIST_OF_FRACTIONS_EXP_EXCEPTION: Nothing to add")
  5.2666 +  | step_add_list_of_fractions_exp (xs)= 
  5.2667 +    let
  5.2668 +        val den_list=termlist2denominators_exp (xs); (* list of denominators *)
  5.2669 +	val (denom,var)=calc_lcm(den_list);      (* common denominator *)
  5.2670 +	val den=factorize_den_exp(#1(den_list),denom,var); (* faktorisierter Nenner !!! *)
  5.2671 +    in
  5.2672 +	com_den_exp(xs,denom,den,var)
  5.2673 +    end;
  5.2674 +
  5.2675 +(* wird aktuell nicht mehr gebraucht, bei rückänderung schon 
  5.2676 +-------------------------------------------------------------
  5.2677 +(* WN0210???SK brauch ma des überhaupt *)
  5.2678 +fun step_add_list_of_fractions2 []=(Free("0",HOLogic.realT),[]:term list)
  5.2679 +  | step_add_list_of_fractions2 [x]=(x,[])
  5.2680 +  | step_add_list_of_fractions2 (xs) = 
  5.2681 +    let
  5.2682 +        val den_list=termlist2denominators (xs); (* list of denominators *)
  5.2683 +	val (denom,var)=calc_lcm(den_list);      (* common denominator *)
  5.2684 +	val den=factorize_den(#1(den_list),denom,var);  (* faktorisierter Nenner !!! *)
  5.2685 +    in
  5.2686 +	(
  5.2687 +	 Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2688 +	 com_den2(xs,denom, poly2term(denom,var)(*den*),var) $
  5.2689 +	 poly2term(denom,var)
  5.2690 +	,
  5.2691 +	[]
  5.2692 +	)
  5.2693 +    end;
  5.2694 +
  5.2695 +(* WN0210???SK brauch ma des überhaupt *)
  5.2696 +fun step_add_list_of_fractions2_exp []=(Free("0",HOLogic.realT),[]:term list)
  5.2697 +  | step_add_list_of_fractions2_exp [x]=(x,[])
  5.2698 +  | step_add_list_of_fractions2_exp (xs) = 
  5.2699 +    let
  5.2700 +        val den_list=termlist2denominators_exp (xs); (* list of denominators *)
  5.2701 +	val (denom,var)=calc_lcm(den_list);      (* common denominator *)
  5.2702 +	val den=factorize_den_exp(#1(den_list),denom,var);  (* faktorisierter Nenner !!! *)
  5.2703 +    in
  5.2704 +	(
  5.2705 +	 Const ("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2706 +	 com_den_exp2(xs,denom, poly2term(denom,var)(*den*),var) $
  5.2707 +	 poly2expanded(denom,var)
  5.2708 +	,
  5.2709 +	[]
  5.2710 +	)
  5.2711 +    end;
  5.2712 +---------------------------------------------- *)
  5.2713 +
  5.2714 +
  5.2715 +(*. converts a term, which contains severel terms seperated by +, into a list of these terms .*)
  5.2716 +fun term2list (t as (Const("HOL.divide",_) $ _ $ _)) = [t]
  5.2717 +  | term2list (t as (Const("Atools.pow",_) $ _ $ _)) = 
  5.2718 +    [Const("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2719 +	  t $ Free("1",HOLogic.realT)
  5.2720 +     ]
  5.2721 +  | term2list (t as (Free(_,_))) = 
  5.2722 +    [Const("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2723 +	  t $  Free("1",HOLogic.realT)
  5.2724 +     ]
  5.2725 +  | term2list (t as (Const("op *",_) $ _ $ _)) = 
  5.2726 +    [Const("HOL.divide",[HOLogic.realT,HOLogic.realT]--->HOLogic.realT) $ 
  5.2727 +	  t $ Free("1",HOLogic.realT)
  5.2728 +     ]
  5.2729 +  | term2list (Const("op +",_) $ t1 $ t2) = term2list(t1) @ term2list(t2)
  5.2730 +  | term2list (Const("op -",_) $ t1 $ t2) = 
  5.2731 +    raise error ("RATIONALS_TERM2LIST_EXCEPTION: - not implemented yet")
  5.2732 +  | term2list _ = raise error ("RATIONALS_TERM2LIST_EXCEPTION: invalid term");
  5.2733 +
  5.2734 +(*.factors out the gcd of nominator and denominator:
  5.2735 +   a/b = (a' * gcd)/(b' * gcd),  a,b,gcd  are poly[2].*)
  5.2736 +fun factout_p_  (thy:theory) t = SOME (step_cancel t,[]:term list); 
  5.2737 +fun factout_ (thy:theory) t = SOME (step_cancel_expanded t,[]:term list); 
  5.2738 +
  5.2739 +(*.cancels a single fraction with normalform [2]
  5.2740 +   resulting in a canceled fraction [2], see factout_ .*)
  5.2741 +fun cancel_p_ (thy:theory) t = (*WN.2.6.03 no rewrite -> NONE !*)
  5.2742 +    (let val (t',asm) = direct_cancel(*_expanded ... corrected MG.21.8.03*) t
  5.2743 +     in if t = t' then NONE else SOME (t',asm) 
  5.2744 +     end) handle _ => NONE;
  5.2745 +(*.the same as above with normalform [3]
  5.2746 +  val cancel_ :
  5.2747 +      theory ->        (*10.02 unused                                    *)
  5.2748 +      term -> 	       (*fraction in normalform [3]                      *)
  5.2749 +      (term * 	       (*fraction in normalform [3]                      *)
  5.2750 +       term list)      (*casual asumptions in normalform [3]             *)
  5.2751 +	  option       (*NONE: the function is not applicable            *).*)
  5.2752 +fun cancel_ (thy:theory) t = SOME (direct_cancel_expanded t) handle _ => NONE;
  5.2753 +
  5.2754 +(*.transforms sums of at least 2 fractions [3] to
  5.2755 +   sums with the least common multiple as nominator.*)
  5.2756 +fun common_nominator_p_ (thy:theory) t =
  5.2757 +((*writeln("### common_nominator_p_ called");*)
  5.2758 +    SOME (step_add_list_of_fractions(term2list(t))) handle _ => NONE
  5.2759 +);
  5.2760 +fun common_nominator_ (thy:theory) t =
  5.2761 +    SOME (step_add_list_of_fractions_exp(term2list(t))) handle _ => NONE;
  5.2762 +
  5.2763 +(*.add 2 or more fractions
  5.2764 +val add_fraction_p_ :
  5.2765 +      theory ->        (*10.02 unused                                    *)
  5.2766 +      term -> 	       (*2 or more fractions with normalform [2]         *)
  5.2767 +      (term * 	       (*one fraction with normalform [2]                *)
  5.2768 +       term list)      (*casual assumptions in normalform [2] WN0210???SK  *)
  5.2769 +	  option       (*NONE: the function is not applicable            *).*)
  5.2770 +fun add_fraction_p_ (thy:theory) t = 
  5.2771 +(writeln("### add_fraction_p_ called");
  5.2772 +    (let val ts = term2list t
  5.2773 +     in if 1 < length ts
  5.2774 +	then SOME (add_list_of_fractions ts)
  5.2775 +	else NONE (*raise error ("RATIONALS_ADD_EXCEPTION: nothing to add")*)
  5.2776 +     end) handle _ => NONE
  5.2777 +);
  5.2778 +(*.same as add_fraction_p_ but with normalform [3].*)
  5.2779 +(*SOME (step_add_list_of_fractions2(term2list(t))); *)
  5.2780 +fun add_fraction_ (thy:theory) t = 
  5.2781 +    if length(term2list(t))>1 
  5.2782 +    then SOME (add_list_of_fractions_exp(term2list(t))) handle _ => NONE
  5.2783 +    else (*raise error ("RATIONALS_ADD_FRACTION_EXCEPTION: nothing to add")*)
  5.2784 +	NONE;
  5.2785 +fun add_fraction_ (thy:theory) t = 
  5.2786 +    (if 1 < length (term2list t)
  5.2787 +     then SOME (add_list_of_fractions_exp (term2list t))
  5.2788 +     else (*raise error ("RATIONALS_ADD_FRACTION_EXCEPTION: nothing to add")*)
  5.2789 +	 NONE) handle _ => NONE;
  5.2790 +
  5.2791 +(*SOME (step_add_list_of_fractions2_exp(term2list(t))); *)
  5.2792 +
  5.2793 +(*. brings the term into a normal form .*)
  5.2794 +fun norm_rational_ (thy:theory) t = 
  5.2795 +    SOME (add_list_of_fractions(term2list(t))) handle _ => NONE; 
  5.2796 +fun norm_expanded_rat_ (thy:theory) t = 
  5.2797 +    SOME (add_list_of_fractions_exp(term2list(t))) handle _ => NONE; 
  5.2798 +
  5.2799 +
  5.2800 +(*.evaluates conditions in calculate_Rational.*)
  5.2801 +(*make local with FIXX@ME result:term *term list*)
  5.2802 +val calc_rat_erls = prep_rls(
  5.2803 +  Rls {id = "calc_rat_erls", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
  5.2804 +	 erls = e_rls, srls = Erls, calc = [], (*asm_thm = [], *)
  5.2805 +	 rules = 
  5.2806 +	 [Calc ("op =",eval_equal "#equal_"),
  5.2807 +	  Calc ("Atools.is'_const",eval_const "#is_const_"),
  5.2808 +	  Thm ("not_true",num_str not_true),
  5.2809 +	  Thm ("not_false",num_str not_false)
  5.2810 +	  ], 
  5.2811 +	 scr = EmptyScr});
  5.2812 +
  5.2813 +
  5.2814 +(*.simplifies expressions with numerals;
  5.2815 +   does NOT rearrange the term by AC-rewriting; thus terms with variables 
  5.2816 +   need to have constants to be commuted together respectively.*)
  5.2817 +val calculate_Rational = prep_rls(
  5.2818 +    merge_rls "calculate_Rational"
  5.2819 +	(Rls {id = "divide", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
  5.2820 +	      erls = calc_rat_erls, srls = Erls, (*asm_thm = [],*) 
  5.2821 +	      calc = [], 
  5.2822 +	      rules = 
  5.2823 +	      [Calc ("HOL.divide"  ,eval_cancel "#divide_"),
  5.2824 +	       
  5.2825 +	       Thm ("sym_real_minus_divide_eq",
  5.2826 +		    num_str (real_minus_divide_eq RS sym)),
  5.2827 +	       (*SYM - ?x / ?y = - (?x / ?y)  may come from subst*)
  5.2828 +	       
  5.2829 +	       Thm ("rat_add",num_str rat_add),
  5.2830 +	       (*"[| a is_const; b is_const; c is_const; d is_const |] ==> \
  5.2831 +		 \"a / c + b / d = (a * d) / (c * d) + (b * c ) / (d * c)"*)
  5.2832 +	       Thm ("rat_add1",num_str rat_add1),
  5.2833 +	       (*"[| a is_const; b is_const; c is_const |] ==> \
  5.2834 +		 \"a / c + b / c = (a + b) / c"*)
  5.2835 +	       Thm ("rat_add2",num_str rat_add2),
  5.2836 +	       (*"[| ?a is_const; ?b is_const; ?c is_const |] ==> \
  5.2837 +		 \?a / ?c + ?b = (?a + ?b * ?c) / ?c"*)
  5.2838 +	       Thm ("rat_add3",num_str rat_add3),
  5.2839 +	       (*"[| a is_const; b is_const; c is_const |] ==> \
  5.2840 +		 \"a + b / c = (a * c) / c + b / c"\
  5.2841 +		 \.... is_const to be omitted here FIXME*)
  5.2842 +	       
  5.2843 +	       Thm ("rat_mult",num_str rat_mult),
  5.2844 +	       (*a / b * (c / d) = a * c / (b * d)*)
  5.2845 +	       Thm ("real_times_divide1_eq",num_str real_times_divide1_eq),
  5.2846 +	       (*?x * (?y / ?z) = ?x * ?y / ?z*)
  5.2847 +	       Thm ("real_times_divide2_eq",num_str real_times_divide2_eq),
  5.2848 +	       (*?y / ?z * ?x = ?y * ?x / ?z*)
  5.2849 +	       
  5.2850 +	       Thm ("real_divide_divide1",num_str real_divide_divide1),
  5.2851 +	       (*"?y ~= 0 ==> ?u / ?v / (?y / ?z) = ?u / ?v * (?z / ?y)"*)
  5.2852 +	       Thm ("real_divide_divide2_eq",num_str real_divide_divide2_eq),
  5.2853 +	       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
  5.2854 +	       
  5.2855 +	       Thm ("rat_power", num_str rat_power),
  5.2856 +	       (*"(?a / ?b) ^^^ ?n = ?a ^^^ ?n / ?b ^^^ ?n"*)
  5.2857 +	       
  5.2858 +	       Thm ("mult_cross",num_str mult_cross),
  5.2859 +	       (*"[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)*)
  5.2860 +	       Thm ("mult_cross1",num_str mult_cross1),
  5.2861 +	       (*"   b ~= 0            ==> (a / b = c    ) = (a     = b * c)*)
  5.2862 +	       Thm ("mult_cross2",num_str mult_cross2)
  5.2863 +	       (*"           d ~= 0    ==> (a     = c / d) = (a * d =     c)*)
  5.2864 +	       ], scr = EmptyScr})
  5.2865 +	calculate_Poly);
  5.2866 +
  5.2867 +
  5.2868 +(*("is_expanded", ("Rational.is'_expanded", eval_is_expanded ""))*)
  5.2869 +fun eval_is_expanded (thmid:string) _ 
  5.2870 +		       (t as (Const("Rational.is'_expanded", _) $ arg)) thy = 
  5.2871 +    if is_expanded arg
  5.2872 +    then SOME (mk_thmid thmid "" 
  5.2873 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
  5.2874 +	       Trueprop $ (mk_equality (t, HOLogic.true_const)))
  5.2875 +    else SOME (mk_thmid thmid "" 
  5.2876 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
  5.2877 +	       Trueprop $ (mk_equality (t, HOLogic.false_const)))
  5.2878 +  | eval_is_expanded _ _ _ _ = NONE; 
  5.2879 +
  5.2880 +val rational_erls = 
  5.2881 +    merge_rls "rational_erls" calculate_Rational 
  5.2882 +	      (append_rls "is_expanded" Atools_erls 
  5.2883 +			  [Calc ("Rational.is'_expanded", eval_is_expanded "")
  5.2884 +			   ]);
  5.2885 +
  5.2886 +
  5.2887 +
  5.2888 +(*.3 'reverse-rewrite-sets' for symbolic computation on rationals:
  5.2889 + =================================================================
  5.2890 + A[2] 'cancel_p': .
  5.2891 + A[3] 'cancel': .
  5.2892 + B[2] 'common_nominator_p': transforms summands in a term [2]
  5.2893 +         to fractions with the (least) common multiple as nominator.
  5.2894 + B[3] 'norm_rational': normalizes arbitrary algebraic terms (without 
  5.2895 +         radicals and transzendental functions) to one canceled fraction,
  5.2896 +	 nominator and denominator in polynomial form.
  5.2897 +
  5.2898 +In order to meet isac's requirements for interactive and stepwise calculation,
  5.2899 +each 'reverse-rewerite-set' consists of an initialization for the interpreter 
  5.2900 +state and of 4 functions, each of which employs rewriting as much as possible.
  5.2901 +The signature of these functions are the same in each 'reverse-rewrite-set' 
  5.2902 +respectively.*)
  5.2903 +
  5.2904 +(* ************************************************************************* *)
  5.2905 +
  5.2906 +
  5.2907 +local(*. cancel_p
  5.2908 +------------------------
  5.2909 +cancels a single fraction consisting of two (uni- or multivariate)
  5.2910 +polynomials WN0609???SK[2] into another such a fraction; examples:
  5.2911 +
  5.2912 +	   a^2 + -1*b^2         a + b
  5.2913 +        -------------------- = ---------
  5.2914 +	a^2 + -2*a*b + b^2     a + -1*b
  5.2915 +
  5.2916 +        a^2    a
  5.2917 +        --- = ---
  5.2918 +         a     1
  5.2919 +
  5.2920 +Remark: the reverse ruleset does _NOT_ work properly with other input !.*)
  5.2921 +(*WN020824 wir werden "uberlegen, wie wir ungeeignete inputs zur"uckweisen*)
  5.2922 +
  5.2923 +val {rules, rew_ord=(_,ro),...} =
  5.2924 +    rep_rls (assoc_rls "make_polynomial");
  5.2925 +(*WN060829 ... make_deriv does not terminate with 1st expl above,
  5.2926 +           see rational.sml --- investigate rulesets for cancel_p ---*)
  5.2927 +val {rules, rew_ord=(_,ro),...} =
  5.2928 +    rep_rls (assoc_rls "rev_rew_p");
  5.2929 +
  5.2930 +val thy = Rational.thy;
  5.2931 +
  5.2932 +(*.init_state = fn : term -> istate
  5.2933 +initialzies the state of the script interpreter. The state is:
  5.2934 +
  5.2935 +type rrlsstate =      (*state for reverse rewriting*)
  5.2936 +     (term *          (*the current formula*)
  5.2937 +      term *          (*the final term*)
  5.2938 +      rule list       (*'reverse rule list' (#)*)
  5.2939 +	    list *    (*may be serveral, eg. in norm_rational*)
  5.2940 +      (rule *         (*Thm (+ Thm generated from Calc) resulting in ...*)
  5.2941 +       (term *        (*... rewrite with ...*)
  5.2942 +	term list))   (*... assumptions*)
  5.2943 +	  list);      (*derivation from given term to normalform
  5.2944 +		       in reverse order with sym_thm;
  5.2945 +                       (#) could be extracted from here by (map #1)*).*)
  5.2946 +(* val {rules, rew_ord=(_,ro),...} =
  5.2947 +       rep_rls (assoc_rls "rev_rew_p")        (*USE ALWAYS, SEE val cancel_p*);
  5.2948 +   val (thy, eval_rls, ro) =(Rational.thy, Atools_erls, ro) (*..val cancel_p*);
  5.2949 +   val t = t;
  5.2950 +   *)
  5.2951 +fun init_state thy eval_rls ro t =
  5.2952 +    let val SOME (t',_) = factout_p_ thy t
  5.2953 +        val SOME (t'',asm) = cancel_p_ thy t
  5.2954 +        val der = reverse_deriv thy eval_rls rules ro NONE t'
  5.2955 +        val der = der @ [(Thm ("real_mult_div_cancel2",
  5.2956 +			       num_str real_mult_div_cancel2),
  5.2957 +			  (t'',asm))]
  5.2958 +        val rs = (distinct_Thm o (map #1)) der
  5.2959 +	val rs = filter_out (eq_Thms ["sym_real_add_zero_left",
  5.2960 +				      "sym_real_mult_0",
  5.2961 +				      "sym_real_mult_1"
  5.2962 +				      (*..insufficient,eg.make_Polynomial*)])rs
  5.2963 +    in (t,t'',[rs(*here only _ONE_ to ease locate_rule*)],der) end;
  5.2964 +
  5.2965 +(*.locate_rule = fn : rule list -> term -> rule
  5.2966 +		      -> (rule * (term * term list) option) list.
  5.2967 +  checks a rule R for being a cancel-rule, and if it is,
  5.2968 +  then return the list of rules (+ the terms they are rewriting to)
  5.2969 +  which need to be applied before R should be applied.
  5.2970 +  precondition: the rule is applicable to the argument-term.
  5.2971 +arguments:
  5.2972 +  rule list: the reverse rule list
  5.2973 +  -> term  : ... to which the rule shall be applied
  5.2974 +  -> rule  : ... to be applied to term
  5.2975 +value:
  5.2976 +  -> (rule           : a rule rewriting to ...
  5.2977 +      * (term        : ... the resulting term ...
  5.2978 +         * term list): ... with the assumptions ( //#0).
  5.2979 +      ) list         : there may be several such rules;
  5.2980 +		       the list is empty, if the rule has nothing to do
  5.2981 +		       with cancelation.*)
  5.2982 +(* val () = ();
  5.2983 +   *)
  5.2984 +fun locate_rule thy eval_rls ro [rs] t r =
  5.2985 +    if (id_of_thm r) mem (map (id_of_thm)) rs
  5.2986 +    then let val ropt =
  5.2987 +		 rewrite_ thy ro eval_rls true (thm_of_thm r) t;
  5.2988 +	 in case ropt of
  5.2989 +		SOME ta => [(r, ta)]
  5.2990 +	      | NONE => (writeln("### locate_rule:  rewrite "^
  5.2991 +				 (id_of_thm r)^" "^(term2str t)^" = NONE");
  5.2992 +			 []) end
  5.2993 +    else (writeln("### locate_rule:  "^(id_of_thm r)^" not mem rrls");[])
  5.2994 +  | locate_rule _ _ _ _ _ _ =
  5.2995 +    raise error ("locate_rule: doesnt match rev-sets in istate");
  5.2996 +
  5.2997 +(*.next_rule = fn : rule list -> term -> rule option
  5.2998 +  for a given term return the next rules to be done for cancelling.
  5.2999 +arguments:
  5.3000 +  rule list     : the reverse rule list
  5.3001 +  term          : the term for which ...
  5.3002 +value:
  5.3003 +  -> rule option: ... this rule is appropriate for cancellation;
  5.3004 +		  there may be no such rule (if the term is canceled already.*)
  5.3005 +(* val thy = Rational.thy;
  5.3006 +   val Rrls {rew_ord=(_,ro),...} = cancel;
  5.3007 +   val ([rs],t) = (rss,f);
  5.3008 +   next_rule thy eval_rls ro [rs] t;(*eval fun next_rule ... before!*)
  5.3009 +
  5.3010 +   val (thy, [rs]) = (Rational.thy, revsets);
  5.3011 +   val Rrls {rew_ord=(_,ro),...} = cancel;
  5.3012 +   nex [rs] t;
  5.3013 +   *)
  5.3014 +fun next_rule thy eval_rls ro [rs] t =
  5.3015 +    let val der = make_deriv thy eval_rls rs ro NONE t;
  5.3016 +    in case der of
  5.3017 +(* val (_,r,_)::_ = der;
  5.3018 +   *)
  5.3019 +	   (_,r,_)::_ => SOME r
  5.3020 +	 | _ => NONE
  5.3021 +    end
  5.3022 +  | next_rule _ _ _ _ _ =
  5.3023 +    raise error ("next_rule: doesnt match rev-sets in istate");
  5.3024 +
  5.3025 +(*.val attach_form = f : rule list -> term -> term
  5.3026 +			 -> (rule * (term * term list)) list
  5.3027 +  checks an input term TI, if it may belong to a current cancellation, by
  5.3028 +  trying to derive it from the given term TG.
  5.3029 +arguments:
  5.3030 +  term   : TG, the last one in the cancellation agreed upon by user + math-eng
  5.3031 +  -> term: TI, the next one input by the user
  5.3032 +value:
  5.3033 +  -> (rule           : the rule to be applied in order to reach TI
  5.3034 +      * (term        : ... obtained by applying the rule ...
  5.3035 +         * term list): ... and the respective assumptions.
  5.3036 +      ) list         : there may be several such rules;
  5.3037 +                       the list is empty, if the users term does not belong
  5.3038 +		       to a cancellation of the term last agreed upon.*)
  5.3039 +fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
  5.3040 +    []:(rule * (term * term list)) list;
  5.3041 +
  5.3042 +in
  5.3043 +
  5.3044 +val cancel_p =
  5.3045 +    Rrls {id = "cancel_p", prepat=[],
  5.3046 +	  rew_ord=("ord_make_polynomial",
  5.3047 +		   ord_make_polynomial false Rational.thy),
  5.3048 +	  erls = rational_erls,
  5.3049 +	  calc = [("PLUS"    ,("op +"        ,eval_binop "#add_")),
  5.3050 +		  ("TIMES"   ,("op *"        ,eval_binop "#mult_")),
  5.3051 +		  ("DIVIDE" ,("HOL.divide"  ,eval_cancel "#divide_")),
  5.3052 +		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
  5.3053 +	  (*asm_thm=[("real_mult_div_cancel2","")],*)
  5.3054 +	  scr=Rfuns {init_state  = init_state thy Atools_erls ro,
  5.3055 +		     normal_form = cancel_p_ thy,
  5.3056 +		     locate_rule = locate_rule thy Atools_erls ro,
  5.3057 +		     next_rule   = next_rule thy Atools_erls ro,
  5.3058 +		     attach_form = attach_form}}
  5.3059 +end;(*local*)
  5.3060 +
  5.3061 +
  5.3062 +local(*.ad (1) 'cancel'
  5.3063 +------------------------------
  5.3064 +cancels a single fraction consisting of two (uni- or multivariate)
  5.3065 +polynomials WN0609???SK[3] into another such a fraction; examples:
  5.3066 +
  5.3067 +	   a^2 - b^2           a + b
  5.3068 +        -------------------- = ---------
  5.3069 +	a^2 - 2*a*b + b^2      a - *b
  5.3070 +
  5.3071 +Remark: the reverse ruleset does _NOT_ work properly with other input !.*)
  5.3072 +(*WN 24.8.02: wir werden "uberlegen, wie wir ungeeignete inputs zur"uckweisen*)
  5.3073 +
  5.3074 +(*
  5.3075 +val SOME (Rls {rules=rules,rew_ord=(_,ro),...}) = 
  5.3076 +    assoc'(!ruleset',"expand_binoms");
  5.3077 +*)
  5.3078 +val {rules=rules,rew_ord=(_,ro),...} =
  5.3079 +    rep_rls (assoc_rls "expand_binoms");
  5.3080 +val thy = Rational.thy;
  5.3081 +
  5.3082 +fun init_state thy eval_rls ro t =
  5.3083 +    let val SOME (t',_) = factout_ thy t;
  5.3084 +        val SOME (t'',asm) = cancel_ thy t;
  5.3085 +        val der = reverse_deriv thy eval_rls rules ro NONE t';
  5.3086 +        val der = der @ [(Thm ("real_mult_div_cancel2",
  5.3087 +			       num_str real_mult_div_cancel2),
  5.3088 +			  (t'',asm))]
  5.3089 +        val rs = map #1 der;
  5.3090 +    in (t,t'',[rs],der) end;
  5.3091 +
  5.3092 +fun locate_rule thy eval_rls ro [rs] t r =
  5.3093 +    if (id_of_thm r) mem (map (id_of_thm)) rs
  5.3094 +    then let val ropt = 
  5.3095 +		 rewrite_ thy ro eval_rls true (thm_of_thm r) t;
  5.3096 +	 in case ropt of
  5.3097 +		SOME ta => [(r, ta)]
  5.3098 +	      | NONE => (writeln("### locate_rule:  rewrite "^
  5.3099 +				 (id_of_thm r)^" "^(term2str t)^" = NONE");
  5.3100 +			 []) end
  5.3101 +    else (writeln("### locate_rule:  "^(id_of_thm r)^" not mem rrls");[])
  5.3102 +  | locate_rule _ _ _ _ _ _ = 
  5.3103 +    raise error ("locate_rule: doesnt match rev-sets in istate");
  5.3104 +
  5.3105 +fun next_rule thy eval_rls ro [rs] t =
  5.3106 +    let val der = make_deriv thy eval_rls rs ro NONE t;
  5.3107 +    in case der of 
  5.3108 +(* val (_,r,_)::_ = der;
  5.3109 +   *)
  5.3110 +	   (_,r,_)::_ => SOME r
  5.3111 +	 | _ => NONE
  5.3112 +    end
  5.3113 +  | next_rule _ _ _ _ _ = 
  5.3114 +    raise error ("next_rule: doesnt match rev-sets in istate");
  5.3115 +
  5.3116 +fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
  5.3117 +    []:(rule * (term * term list)) list;
  5.3118 +
  5.3119 +val pat = (term_of o the o (parse thy)) "?r/?s";
  5.3120 +val pre1 = (term_of o the o (parse thy)) "?r is_expanded";
  5.3121 +val pre2 = (term_of o the o (parse thy)) "?s is_expanded";
  5.3122 +val prepat = [([pre1, pre2], pat)];
  5.3123 +
  5.3124 +in
  5.3125 +
  5.3126 +
  5.3127 +val cancel = 
  5.3128 +    Rrls {id = "cancel", prepat=prepat,
  5.3129 +	  rew_ord=("ord_make_polynomial",
  5.3130 +		   ord_make_polynomial false Rational.thy),
  5.3131 +	  erls = rational_erls, 
  5.3132 +	  calc = [("PLUS"    ,("op +"        ,eval_binop "#add_")),
  5.3133 +		  ("TIMES"   ,("op *"        ,eval_binop "#mult_")),
  5.3134 +		  ("DIVIDE" ,("HOL.divide"  ,eval_cancel "#divide_")),
  5.3135 +		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
  5.3136 +	  scr=Rfuns {init_state  = init_state thy Atools_erls ro,
  5.3137 +		     normal_form = cancel_ thy, 
  5.3138 +		     locate_rule = locate_rule thy Atools_erls ro,
  5.3139 +		     next_rule   = next_rule thy Atools_erls ro,
  5.3140 +		     attach_form = attach_form}}
  5.3141 +end;(*local*)
  5.3142 +
  5.3143 +
  5.3144 +
  5.3145 +local(*.ad [2] 'common_nominator_p'
  5.3146 +---------------------------------
  5.3147 +FIXME Beschreibung .*)
  5.3148 +
  5.3149 +
  5.3150 +val {rules=rules,rew_ord=(_,ro),...} =
  5.3151 +    rep_rls (assoc_rls "make_polynomial");
  5.3152 +(*WN060829 ... make_deriv does not terminate with 1st expl above,
  5.3153 +           see rational.sml --- investigate rulesets for cancel_p ---*)
  5.3154 +val {rules, rew_ord=(_,ro),...} =
  5.3155 +    rep_rls (assoc_rls "rev_rew_p");
  5.3156 +val thy = Rational.thy;
  5.3157 +
  5.3158 +
  5.3159 +(*.common_nominator_p_ = fn : theory -> term -> (term * term list) option
  5.3160 +  as defined above*)
  5.3161 +
  5.3162 +(*.init_state = fn : term -> istate
  5.3163 +initialzies the state of the interactive interpreter. The state is:
  5.3164 +
  5.3165 +type rrlsstate =      (*state for reverse rewriting*)
  5.3166 +     (term *          (*the current formula*)
  5.3167 +      term *          (*the final term*)
  5.3168 +      rule list       (*'reverse rule list' (#)*)
  5.3169 +	    list *    (*may be serveral, eg. in norm_rational*)
  5.3170 +      (rule *         (*Thm (+ Thm generated from Calc) resulting in ...*)
  5.3171 +       (term *        (*... rewrite with ...*)
  5.3172 +	term list))   (*... assumptions*)
  5.3173 +	  list);      (*derivation from given term to normalform
  5.3174 +		       in reverse order with sym_thm;
  5.3175 +                       (#) could be extracted from here by (map #1)*).*)
  5.3176 +fun init_state thy eval_rls ro t =
  5.3177 +    let val SOME (t',_) = common_nominator_p_ thy t;
  5.3178 +        val SOME (t'',asm) = add_fraction_p_ thy t;
  5.3179 +        val der = reverse_deriv thy eval_rls rules ro NONE t';
  5.3180 +        val der = der @ [(Thm ("real_mult_div_cancel2",
  5.3181 +			       num_str real_mult_div_cancel2),
  5.3182 +			  (t'',asm))]
  5.3183 +        val rs = (distinct_Thm o (map #1)) der;
  5.3184 +	val rs = filter_out (eq_Thms ["sym_real_add_zero_left",
  5.3185 +				      "sym_real_mult_0",
  5.3186 +				      "sym_real_mult_1"]) rs;
  5.3187 +    in (t,t'',[rs(*here only _ONE_*)],der) end;
  5.3188 +
  5.3189 +(* use"knowledge/Rational.ML";
  5.3190 +   *)
  5.3191 +
  5.3192 +(*.locate_rule = fn : rule list -> term -> rule
  5.3193 +		      -> (rule * (term * term list) option) list.
  5.3194 +  checks a rule R for being a cancel-rule, and if it is,
  5.3195 +  then return the list of rules (+ the terms they are rewriting to)
  5.3196 +  which need to be applied before R should be applied.
  5.3197 +  precondition: the rule is applicable to the argument-term.
  5.3198 +arguments:
  5.3199 +  rule list: the reverse rule list
  5.3200 +  -> term  : ... to which the rule shall be applied
  5.3201 +  -> rule  : ... to be applied to term
  5.3202 +value:
  5.3203 +  -> (rule           : a rule rewriting to ...
  5.3204 +      * (term        : ... the resulting term ...
  5.3205 +         * term list): ... with the assumptions ( //#0).
  5.3206 +      ) list         : there may be several such rules;
  5.3207 +		       the list is empty, if the rule has nothing to do
  5.3208 +		       with cancelation.*)
  5.3209 +(* val () = ();
  5.3210 +   *)
  5.3211 +fun locate_rule thy eval_rls ro [rs] t r =
  5.3212 +    if (id_of_thm r) mem (map (id_of_thm)) rs
  5.3213 +    then let val ropt =
  5.3214 +		 rewrite_ thy ro eval_rls true (thm_of_thm r) t;
  5.3215 +	 in case ropt of
  5.3216 +		SOME ta => [(r, ta)]
  5.3217 +	      | NONE => (writeln("### locate_rule:  rewrite "^
  5.3218 +				 (id_of_thm r)^" "^(term2str t)^" = NONE");
  5.3219 +			 []) end
  5.3220 +    else (writeln("### locate_rule:  "^(id_of_thm r)^" not mem rrls");[])
  5.3221 +  | locate_rule _ _ _ _ _ _ =
  5.3222 +    raise error ("locate_rule: doesnt match rev-sets in istate");
  5.3223 +
  5.3224 +(*.next_rule = fn : rule list -> term -> rule option
  5.3225 +  for a given term return the next rules to be done for cancelling.
  5.3226 +arguments:
  5.3227 +  rule list     : the reverse rule list
  5.3228 +  term          : the term for which ...
  5.3229 +value:
  5.3230 +  -> rule option: ... this rule is appropriate for cancellation;
  5.3231 +		  there may be no such rule (if the term is canceled already.*)
  5.3232 +(* val thy = Rational.thy;
  5.3233 +   val Rrls {rew_ord=(_,ro),...} = cancel;
  5.3234 +   val ([rs],t) = (rss,f);
  5.3235 +   next_rule thy eval_rls ro [rs] t;(*eval fun next_rule ... before!*)
  5.3236 +
  5.3237 +   val (thy, [rs]) = (Rational.thy, revsets);
  5.3238 +   val Rrls {rew_ord=(_,ro),...} = cancel;
  5.3239 +   nex [rs] t;
  5.3240 +   *)
  5.3241 +fun next_rule thy eval_rls ro [rs] t =
  5.3242 +    let val der = make_deriv thy eval_rls rs ro NONE t;
  5.3243 +    in case der of
  5.3244 +(* val (_,r,_)::_ = der;
  5.3245 +   *)
  5.3246 +	   (_,r,_)::_ => SOME r
  5.3247 +	 | _ => NONE
  5.3248 +    end
  5.3249 +  | next_rule _ _ _ _ _ =
  5.3250 +    raise error ("next_rule: doesnt match rev-sets in istate");
  5.3251 +
  5.3252 +(*.val attach_form = f : rule list -> term -> term
  5.3253 +			 -> (rule * (term * term list)) list
  5.3254 +  checks an input term TI, if it may belong to a current cancellation, by
  5.3255 +  trying to derive it from the given term TG.
  5.3256 +arguments:
  5.3257 +  term   : TG, the last one in the cancellation agreed upon by user + math-eng
  5.3258 +  -> term: TI, the next one input by the user
  5.3259 +value:
  5.3260 +  -> (rule           : the rule to be applied in order to reach TI
  5.3261 +      * (term        : ... obtained by applying the rule ...
  5.3262 +         * term list): ... and the respective assumptions.
  5.3263 +      ) list         : there may be several such rules;
  5.3264 +                       the list is empty, if the users term does not belong
  5.3265 +		       to a cancellation of the term last agreed upon.*)
  5.3266 +fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
  5.3267 +    []:(rule * (term * term list)) list;
  5.3268 +
  5.3269 +val pat0 = (term_of o the o (parse thy)) "?r/?s+?u/?v";
  5.3270 +val pat1 = (term_of o the o (parse thy)) "?r/?s+?u   ";
  5.3271 +val pat2 = (term_of o the o (parse thy)) "?r   +?u/?v";
  5.3272 +val prepat = [([HOLogic.true_const], pat0),
  5.3273 +	      ([HOLogic.true_const], pat1),
  5.3274 +	      ([HOLogic.true_const], pat2)];
  5.3275 +
  5.3276 +in
  5.3277 +
  5.3278 +(*11.02 schnelle L"osung f"ur RL: Bruch auch gek"urzt;
  5.3279 +  besser w"are: auf 1 gemeinsamen Bruchstrich, Nenner und Z"ahler unvereinfacht
  5.3280 +  dh. wie common_nominator_p_, aber auf 1 Bruchstrich*)
  5.3281 +val common_nominator_p =
  5.3282 +    Rrls {id = "common_nominator_p", prepat=prepat,
  5.3283 +	  rew_ord=("ord_make_polynomial",
  5.3284 +		   ord_make_polynomial false Rational.thy),
  5.3285 +	  erls = rational_erls,
  5.3286 +	  calc = [("PLUS"    ,("op +"        ,eval_binop "#add_")),
  5.3287 +		  ("TIMES"   ,("op *"        ,eval_binop "#mult_")),
  5.3288 +		  ("DIVIDE" ,("HOL.divide"  ,eval_cancel "#divide_")),
  5.3289 +		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
  5.3290 +	  scr=Rfuns {init_state  = init_state thy Atools_erls ro,
  5.3291 +		     normal_form = add_fraction_p_ thy,(*FIXME.WN0211*)
  5.3292 +		     locate_rule = locate_rule thy Atools_erls ro,
  5.3293 +		     next_rule   = next_rule thy Atools_erls ro,
  5.3294 +		     attach_form = attach_form}}
  5.3295 +end;(*local*)
  5.3296 +
  5.3297 +
  5.3298 +local(*.ad [2] 'common_nominator'
  5.3299 +---------------------------------
  5.3300 +FIXME Beschreibung .*)
  5.3301 +
  5.3302 +
  5.3303 +val {rules=rules,rew_ord=(_,ro),...} =
  5.3304 +    rep_rls (assoc_rls "make_polynomial");
  5.3305 +val thy = Rational.thy;
  5.3306 +
  5.3307 +
  5.3308 +(*.common_nominator_ = fn : theory -> term -> (term * term list) option
  5.3309 +  as defined above*)
  5.3310 +
  5.3311 +(*.init_state = fn : term -> istate
  5.3312 +initialzies the state of the interactive interpreter. The state is:
  5.3313 +
  5.3314 +type rrlsstate =      (*state for reverse rewriting*)
  5.3315 +     (term *          (*the current formula*)
  5.3316 +      term *          (*the final term*)
  5.3317 +      rule list       (*'reverse rule list' (#)*)
  5.3318 +	    list *    (*may be serveral, eg. in norm_rational*)
  5.3319 +      (rule *         (*Thm (+ Thm generated from Calc) resulting in ...*)
  5.3320 +       (term *        (*... rewrite with ...*)
  5.3321 +	term list))   (*... assumptions*)
  5.3322 +	  list);      (*derivation from given term to normalform
  5.3323 +		       in reverse order with sym_thm;
  5.3324 +                       (#) could be extracted from here by (map #1)*).*)
  5.3325 +fun init_state thy eval_rls ro t =
  5.3326 +    let val SOME (t',_) = common_nominator_ thy t;
  5.3327 +        val SOME (t'',asm) = add_fraction_ thy t;
  5.3328 +        val der = reverse_deriv thy eval_rls rules ro NONE t';
  5.3329 +        val der = der @ [(Thm ("real_mult_div_cancel2",
  5.3330 +			       num_str real_mult_div_cancel2),
  5.3331 +			  (t'',asm))]
  5.3332 +        val rs = (distinct_Thm o (map #1)) der;
  5.3333 +	val rs = filter_out (eq_Thms ["sym_real_add_zero_left",
  5.3334 +				      "sym_real_mult_0",
  5.3335 +				      "sym_real_mult_1"]) rs;
  5.3336 +    in (t,t'',[rs(*here only _ONE_*)],der) end;
  5.3337 +
  5.3338 +(* use"knowledge/Rational.ML";
  5.3339 +   *)
  5.3340 +
  5.3341 +(*.locate_rule = fn : rule list -> term -> rule
  5.3342 +		      -> (rule * (term * term list) option) list.
  5.3343 +  checks a rule R for being a cancel-rule, and if it is,
  5.3344 +  then return the list of rules (+ the terms they are rewriting to)
  5.3345 +  which need to be applied before R should be applied.
  5.3346 +  precondition: the rule is applicable to the argument-term.
  5.3347 +arguments:
  5.3348 +  rule list: the reverse rule list
  5.3349 +  -> term  : ... to which the rule shall be applied
  5.3350 +  -> rule  : ... to be applied to term
  5.3351 +value:
  5.3352 +  -> (rule           : a rule rewriting to ...
  5.3353 +      * (term        : ... the resulting term ...
  5.3354 +         * term list): ... with the assumptions ( //#0).
  5.3355 +      ) list         : there may be several such rules;
  5.3356 +		       the list is empty, if the rule has nothing to do
  5.3357 +		       with cancelation.*)
  5.3358 +(* val () = ();
  5.3359 +   *)
  5.3360 +fun locate_rule thy eval_rls ro [rs] t r =
  5.3361 +    if (id_of_thm r) mem (map (id_of_thm)) rs
  5.3362 +    then let val ropt =
  5.3363 +		 rewrite_ thy ro eval_rls true (thm_of_thm r) t;
  5.3364 +	 in case ropt of
  5.3365 +		SOME ta => [(r, ta)]
  5.3366 +	      | NONE => (writeln("### locate_rule:  rewrite "^
  5.3367 +				 (id_of_thm r)^" "^(term2str t)^" = NONE");
  5.3368 +			 []) end
  5.3369 +    else (writeln("### locate_rule:  "^(id_of_thm r)^" not mem rrls");[])
  5.3370 +  | locate_rule _ _ _ _ _ _ =
  5.3371 +    raise error ("locate_rule: doesnt match rev-sets in istate");
  5.3372 +
  5.3373 +(*.next_rule = fn : rule list -> term -> rule option
  5.3374 +  for a given term return the next rules to be done for cancelling.
  5.3375 +arguments:
  5.3376 +  rule list     : the reverse rule list
  5.3377 +  term          : the term for which ...
  5.3378 +value:
  5.3379 +  -> rule option: ... this rule is appropriate for cancellation;
  5.3380 +		  there may be no such rule (if the term is canceled already.*)
  5.3381 +(* val thy = Rational.thy;
  5.3382 +   val Rrls {rew_ord=(_,ro),...} = cancel;
  5.3383 +   val ([rs],t) = (rss,f);
  5.3384 +   next_rule thy eval_rls ro [rs] t;(*eval fun next_rule ... before!*)
  5.3385 +
  5.3386 +   val (thy, [rs]) = (Rational.thy, revsets);
  5.3387 +   val Rrls {rew_ord=(_,ro),...} = cancel_p;
  5.3388 +   nex [rs] t;
  5.3389 +   *)
  5.3390 +fun next_rule thy eval_rls ro [rs] t =
  5.3391 +    let val der = make_deriv thy eval_rls rs ro NONE t;
  5.3392 +    in case der of
  5.3393 +(* val (_,r,_)::_ = der;
  5.3394 +   *)
  5.3395 +	   (_,r,_)::_ => SOME r
  5.3396 +	 | _ => NONE
  5.3397 +    end
  5.3398 +  | next_rule _ _ _ _ _ =
  5.3399 +    raise error ("next_rule: doesnt match rev-sets in istate");
  5.3400 +
  5.3401 +(*.val attach_form = f : rule list -> term -> term
  5.3402 +			 -> (rule * (term * term list)) list
  5.3403 +  checks an input term TI, if it may belong to a current cancellation, by
  5.3404 +  trying to derive it from the given term TG.
  5.3405 +arguments:
  5.3406 +  term   : TG, the last one in the cancellation agreed upon by user + math-eng
  5.3407 +  -> term: TI, the next one input by the user
  5.3408 +value:
  5.3409 +  -> (rule           : the rule to be applied in order to reach TI
  5.3410 +      * (term        : ... obtained by applying the rule ...
  5.3411 +         * term list): ... and the respective assumptions.
  5.3412 +      ) list         : there may be several such rules;
  5.3413 +                       the list is empty, if the users term does not belong
  5.3414 +		       to a cancellation of the term last agreed upon.*)
  5.3415 +fun attach_form (_:rule list list) (_:term) (_:term) = (*still missing*)
  5.3416 +    []:(rule * (term * term list)) list;
  5.3417 +
  5.3418 +val pat0 =  (term_of o the o (parse thy)) "?r/?s+?u/?v";
  5.3419 +val pat01 = (term_of o the o (parse thy)) "?r/?s-?u/?v";
  5.3420 +val pat1 =  (term_of o the o (parse thy)) "?r/?s+?u   ";
  5.3421 +val pat11 = (term_of o the o (parse thy)) "?r/?s-?u   ";
  5.3422 +val pat2 =  (term_of o the o (parse thy)) "?r   +?u/?v";
  5.3423 +val pat21 = (term_of o the o (parse thy)) "?r   -?u/?v";
  5.3424 +val prepat = [([HOLogic.true_const], pat0),
  5.3425 +	      ([HOLogic.true_const], pat01),
  5.3426 +	      ([HOLogic.true_const], pat1),
  5.3427 +	      ([HOLogic.true_const], pat11),
  5.3428 +	      ([HOLogic.true_const], pat2),
  5.3429 +	      ([HOLogic.true_const], pat21)];
  5.3430 +
  5.3431 +
  5.3432 +in
  5.3433 +
  5.3434 +val common_nominator =
  5.3435 +    Rrls {id = "common_nominator", prepat=prepat,
  5.3436 +	  rew_ord=("ord_make_polynomial",
  5.3437 +		   ord_make_polynomial false Rational.thy),
  5.3438 +	  erls = rational_erls,
  5.3439 +	  calc = [("PLUS"    ,("op +"        ,eval_binop "#add_")),
  5.3440 +		  ("TIMES"   ,("op *"        ,eval_binop "#mult_")),
  5.3441 +		  ("DIVIDE" ,("HOL.divide"  ,eval_cancel "#divide_")),
  5.3442 +		  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))],
  5.3443 +	  (*asm_thm=[("real_mult_div_cancel2","")],*)
  5.3444 +	  scr=Rfuns {init_state  = init_state thy Atools_erls ro,
  5.3445 +		     normal_form = add_fraction_ (*NOT common_nominator_*) thy,
  5.3446 +		     locate_rule = locate_rule thy Atools_erls ro,
  5.3447 +		     next_rule   = next_rule thy Atools_erls ro,
  5.3448 +		     attach_form = attach_form}}
  5.3449 +
  5.3450 +end;(*local*)
  5.3451 +
  5.3452 +
  5.3453 +(*##*)
  5.3454 +end;(*struct*)
  5.3455 +
  5.3456 +open RationalI;
  5.3457 +(*##*)
  5.3458 +
  5.3459 +(*.the expression contains + - * ^ / only ?.*)
  5.3460 +fun is_ratpolyexp (Free _) = true
  5.3461 +  | is_ratpolyexp (Const ("op +",_) $ Free _ $ Free _) = true
  5.3462 +  | is_ratpolyexp (Const ("op -",_) $ Free _ $ Free _) = true
  5.3463 +  | is_ratpolyexp (Const ("op *",_) $ Free _ $ Free _) = true
  5.3464 +  | is_ratpolyexp (Const ("Atools.pow",_) $ Free _ $ Free _) = true
  5.3465 +  | is_ratpolyexp (Const ("HOL.divide",_) $ Free _ $ Free _) = true
  5.3466 +  | is_ratpolyexp (Const ("op +",_) $ t1 $ t2) = 
  5.3467 +               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
  5.3468 +  | is_ratpolyexp (Const ("op -",_) $ t1 $ t2) = 
  5.3469 +               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
  5.3470 +  | is_ratpolyexp (Const ("op *",_) $ t1 $ t2) = 
  5.3471 +               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
  5.3472 +  | is_ratpolyexp (Const ("Atools.pow",_) $ t1 $ t2) = 
  5.3473 +               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
  5.3474 +  | is_ratpolyexp (Const ("HOL.divide",_) $ t1 $ t2) = 
  5.3475 +               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
  5.3476 +  | is_ratpolyexp _ = false;
  5.3477 +
  5.3478 +(*("is_ratpolyexp", ("Rational.is'_ratpolyexp", eval_is_ratpolyexp ""))*)
  5.3479 +fun eval_is_ratpolyexp (thmid:string) _ 
  5.3480 +		       (t as (Const("Rational.is'_ratpolyexp", _) $ arg)) thy =
  5.3481 +    if is_ratpolyexp arg
  5.3482 +    then SOME (mk_thmid thmid "" 
  5.3483 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
  5.3484 +	       Trueprop $ (mk_equality (t, HOLogic.true_const)))
  5.3485 +    else SOME (mk_thmid thmid "" 
  5.3486 +			((Syntax.string_of_term (thy2ctxt thy)) arg) "", 
  5.3487 +	       Trueprop $ (mk_equality (t, HOLogic.false_const)))
  5.3488 +  | eval_is_ratpolyexp _ _ _ _ = NONE; 
  5.3489 +
  5.3490 +
  5.3491 +
  5.3492 +(*-------------------18.3.03 --> struct <-----------vvv--*)
  5.3493 +val add_fractions_p = common_nominator_p; (*FIXXXME:eilig f"ur norm_Rational*)
  5.3494 +
  5.3495 +(*.discard binary minus, shift unary minus into -1*; 
  5.3496 +   unary minus before numerals are put into the numeral by parsing;
  5.3497 +   contains absolute minimum of thms for context in norm_Rational .*)
  5.3498 +val discard_minus = prep_rls(
  5.3499 +  Rls {id = "discard_minus", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
  5.3500 +      erls = e_rls, srls = Erls, calc = [], (*asm_thm = [],*)
  5.3501 +      rules = [Thm ("real_diff_minus", num_str real_diff_minus),
  5.3502 +	       (*"a - b = a + -1 * b"*)
  5.3503 +	       Thm ("sym_real_mult_minus1",num_str (real_mult_minus1 RS sym))
  5.3504 +	       (*- ?z = "-1 * ?z"*)
  5.3505 +	       ],
  5.3506 +      scr = Script ((term_of o the o (parse thy)) 
  5.3507 +      "empty_script")
  5.3508 +      }):rls;
  5.3509 +(*erls for calculate_Rational; make local with FIXX@ME result:term *term list*)
  5.3510 +val powers_erls = prep_rls(
  5.3511 +  Rls {id = "powers_erls", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
  5.3512 +      erls = e_rls, srls = Erls, calc = [], (*asm_thm = [],*)
  5.3513 +      rules = [Calc ("Atools.is'_atom",eval_is_atom "#is_atom_"),
  5.3514 +	       Calc ("Atools.is'_even",eval_is_even "#is_even_"),
  5.3515 +	       Calc ("op <",eval_equ "#less_"),
  5.3516 +	       Thm ("not_false", not_false),
  5.3517 +	       Thm ("not_true", not_true),
  5.3518 +	       Calc ("op +",eval_binop "#add_")
  5.3519 +	       ],
  5.3520 +      scr = Script ((term_of o the o (parse thy)) 
  5.3521 +      "empty_script")
  5.3522 +      }:rls);
  5.3523 +(*.all powers over + distributed; atoms over * collected, other distributed
  5.3524 +   contains absolute minimum of thms for context in norm_Rational .*)
  5.3525 +val powers = prep_rls(
  5.3526 +  Rls {id = "powers", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
  5.3527 +      erls = powers_erls, srls = Erls, calc = [], (*asm_thm = [],*)
  5.3528 +      rules = [Thm ("realpow_multI", num_str realpow_multI),
  5.3529 +	       (*"(r * s) ^^^ n = r ^^^ n * s ^^^ n"*)
  5.3530 +	       Thm ("realpow_pow",num_str realpow_pow),
  5.3531 +	       (*"(a ^^^ b) ^^^ c = a ^^^ (b * c)"*)
  5.3532 +	       Thm ("realpow_oneI",num_str realpow_oneI),
  5.3533 +	       (*"r ^^^ 1 = r"*)
  5.3534 +	       Thm ("realpow_minus_even",num_str realpow_minus_even),
  5.3535 +	       (*"n is_even ==> (- r) ^^^ n = r ^^^ n" ?-->discard_minus?*)
  5.3536 +	       Thm ("realpow_minus_odd",num_str realpow_minus_odd),
  5.3537 +	       (*"Not (n is_even) ==> (- r) ^^^ n = -1 * r ^^^ n"*)
  5.3538 +	       
  5.3539 +	       (*----- collect atoms over * -----*)
  5.3540 +	       Thm ("realpow_two_atom",num_str realpow_two_atom),	
  5.3541 +	       (*"r is_atom ==> r * r = r ^^^ 2"*)
  5.3542 +	       Thm ("realpow_plus_1",num_str realpow_plus_1),		
  5.3543 +	       (*"r is_atom ==> r * r ^^^ n = r ^^^ (n + 1)"*)
  5.3544 +	       Thm ("realpow_addI_atom",num_str realpow_addI_atom),
  5.3545 +	       (*"r is_atom ==> r ^^^ n * r ^^^ m = r ^^^ (n + m)"*)
  5.3546 +
  5.3547 +	       (*----- distribute none-atoms -----*)
  5.3548 +	       Thm ("realpow_def_atom",num_str realpow_def_atom),
  5.3549 +	       (*"[| 1 < n; not(r is_atom) |]==>r ^^^ n = r * r ^^^ (n + -1)"*)
  5.3550 +	       Thm ("realpow_eq_oneI",num_str realpow_eq_oneI),
  5.3551 +	       (*"1 ^^^ n = 1"*)
  5.3552 +	       Calc ("op +",eval_binop "#add_")
  5.3553 +	       ],
  5.3554 +      scr = Script ((term_of o the o (parse thy)) 
  5.3555 +      "empty_script")
  5.3556 +      }:rls);
  5.3557 +(*.contains absolute minimum of thms for context in norm_Rational.*)
  5.3558 +val rat_mult_divide = prep_rls(
  5.3559 +  Rls {id = "rat_mult_divide", preconds = [], 
  5.3560 +       rew_ord = ("dummy_ord",dummy_ord), 
  5.3561 +      erls = e_rls, srls = Erls, calc = [], (*asm_thm = [],*)
  5.3562 +      rules = [Thm ("rat_mult",num_str rat_mult),
  5.3563 +	       (*(1)"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
  5.3564 +	       Thm ("real_times_divide1_eq",num_str real_times_divide1_eq),
  5.3565 +	       (*(2)"?a * (?c / ?d) = ?a * ?c / ?d" must be [2],
  5.3566 +	       otherwise inv.to a / b / c = ...*)
  5.3567 +	       Thm ("real_times_divide2_eq",num_str real_times_divide2_eq),
  5.3568 +	       (*"?a / ?b * ?c = ?a * ?c / ?b" order weights x^^^n too much
  5.3569 +		     and does not commute a / b * c ^^^ 2 !*)
  5.3570 +	       
  5.3571 +	       Thm ("real_divide_divide1_eq", real_divide_divide1_eq),
  5.3572 +	       (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
  5.3573 +	       Thm ("real_divide_divide2_eq", real_divide_divide2_eq),
  5.3574 +	       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
  5.3575 +	       Calc ("HOL.divide"  ,eval_cancel "#divide_")
  5.3576 +	       ],
  5.3577 +      scr = Script ((term_of o the o (parse thy)) "empty_script")
  5.3578 +      }:rls);
  5.3579 +(*.contains absolute minimum of thms for context in norm_Rational.*)
  5.3580 +val reduce_0_1_2 = prep_rls(
  5.3581 +  Rls{id = "reduce_0_1_2", preconds = [], rew_ord = ("dummy_ord", dummy_ord),
  5.3582 +      erls = e_rls,srls = Erls,calc = [],(*asm_thm = [],*)
  5.3583 +      rules = [(*Thm ("real_divide_1",num_str real_divide_1),
  5.3584 +		 "?x / 1 = ?x" unnecess.for normalform*)
  5.3585 +	       Thm ("real_mult_1",num_str real_mult_1),                 
  5.3586 +	       (*"1 * z = z"*)
  5.3587 +	       (*Thm ("real_mult_minus1",num_str real_mult_minus1),
  5.3588 +	       "-1 * z = - z"*)
  5.3589 +	       (*Thm ("real_minus_mult_cancel",num_str real_minus_mult_cancel),
  5.3590 +	       "- ?x * - ?y = ?x * ?y"*)
  5.3591 +
  5.3592 +	       Thm ("real_mult_0",num_str real_mult_0),        
  5.3593 +	       (*"0 * z = 0"*)
  5.3594 +	       Thm ("real_add_zero_left",num_str real_add_zero_left),
  5.3595 +	       (*"0 + z = z"*)
  5.3596 +	       (*Thm ("real_add_minus",num_str real_add_minus),
  5.3597 +	       "?z + - ?z = 0"*)
  5.3598 +
  5.3599 +	       Thm ("sym_real_mult_2",num_str (real_mult_2 RS sym)),	
  5.3600 +	       (*"z1 + z1 = 2 * z1"*)
  5.3601 +	       Thm ("real_mult_2_assoc",num_str real_mult_2_assoc),
  5.3602 +	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
  5.3603 +
  5.3604 +	       Thm ("real_0_divide",num_str real_0_divide)
  5.3605 +	       (*"0 / ?x = 0"*)
  5.3606 +	       ], scr = EmptyScr}:rls);
  5.3607 +
  5.3608 +(*erls for calculate_Rational; 
  5.3609 +  make local with FIXX@ME result:term *term list WN0609???SKMG*)
  5.3610 +val norm_rat_erls = prep_rls(
  5.3611 +  Rls {id = "norm_rat_erls", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
  5.3612 +      erls = e_rls, srls = Erls, calc = [], (*asm_thm = [],*)
  5.3613 +      rules = [Calc ("Atools.is'_const",eval_const "#is_const_")
  5.3614 +	       ],
  5.3615 +      scr = Script ((term_of o the o (parse thy)) 
  5.3616 +      "empty_script")
  5.3617 +      }:rls);
  5.3618 +(*.consists of rls containing the absolute minimum of thms.*)
  5.3619 +(*040209: this version has been used by RL for his equations,
  5.3620 +which is now replaced by MGs version below
  5.3621 +vvv OLD VERSION !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!*)
  5.3622 +val norm_Rational = prep_rls(
  5.3623 +  Rls {id = "norm_Rational", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
  5.3624 +      erls = norm_rat_erls, srls = Erls, calc = [], (*asm_thm = [],*)
  5.3625 +      rules = [(*sequence given by operator precedence*)
  5.3626 +	       Rls_ discard_minus,
  5.3627 +	       Rls_ powers,
  5.3628 +	       Rls_ rat_mult_divide,
  5.3629 +	       Rls_ expand,
  5.3630 +	       Rls_ reduce_0_1_2,
  5.3631 +	       (*^^^^^^^^^ from RL -- not the latest one vvvvvvvvv*)
  5.3632 +	       Rls_ order_add_mult,
  5.3633 +	       Rls_ collect_numerals,
  5.3634 +	       Rls_ add_fractions_p,
  5.3635 +	       Rls_ cancel_p
  5.3636 +	       ],
  5.3637 +      scr = Script ((term_of o the o (parse thy)) 
  5.3638 +      "empty_script")
  5.3639 +      }:rls);
  5.3640 +val norm_Rational_parenthesized = prep_rls(
  5.3641 +  Seq {id = "norm_Rational_parenthesized", preconds = []:term list, 
  5.3642 +       rew_ord = ("dummy_ord", dummy_ord),
  5.3643 +      erls = Atools_erls, srls = Erls,
  5.3644 +      calc = [], (*asm_thm = [],*)
  5.3645 +      rules = [Rls_  norm_Rational, (*from RL -- not the latest one*)
  5.3646 +	       Rls_ discard_parentheses
  5.3647 +	       ],
  5.3648 +      scr = EmptyScr
  5.3649 +      }:rls);      
  5.3650 +
  5.3651 +
  5.3652 +(*-------------------18.3.03 --> struct <-----------^^^--*)
  5.3653 +
  5.3654 +
  5.3655 +
  5.3656 +theory' := overwritel (!theory', [("Rational.thy",Rational.thy)]);
  5.3657 +
  5.3658 +
  5.3659 +(*WN030318???SK: simplifies all but cancel and common_nominator*)
  5.3660 +val simplify_rational = 
  5.3661 +    merge_rls "simplify_rational" expand_binoms
  5.3662 +    (append_rls "divide" calculate_Rational
  5.3663 +		[Thm ("real_divide_1",num_str real_divide_1),
  5.3664 +		 (*"?x / 1 = ?x"*)
  5.3665 +		 Thm ("rat_mult",num_str rat_mult),
  5.3666 +		 (*(1)"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
  5.3667 +		 Thm ("real_times_divide1_eq",num_str real_times_divide1_eq),
  5.3668 +		 (*(2)"?a * (?c / ?d) = ?a * ?c / ?d" must be [2],
  5.3669 +		 otherwise inv.to a / b / c = ...*)
  5.3670 +		 Thm ("real_times_divide2_eq",num_str real_times_divide2_eq),
  5.3671 +		 (*"?a / ?b * ?c = ?a * ?c / ?b"*)
  5.3672 +		 Thm ("add_minus",num_str add_minus),
  5.3673 +		 (*"?a + ?b - ?b = ?a"*)
  5.3674 +		 Thm ("add_minus1",num_str add_minus1),
  5.3675 +		 (*"?a - ?b + ?b = ?a"*)
  5.3676 +		 Thm ("real_divide_minus1",num_str real_divide_minus1)
  5.3677 +		 (*"?x / -1 = - ?x"*)
  5.3678 +(*
  5.3679 +,
  5.3680 +		 Thm ("",num_str )
  5.3681 +*)
  5.3682 +		 ]);
  5.3683 +
  5.3684 +(*---------vvv-------------MG ab 1.07.2003--------------vvv-----------*)
  5.3685 +
  5.3686 +(* ------------------------------------------------------------------ *)
  5.3687 +(*                  Simplifier für beliebige Buchterme                *) 
  5.3688 +(* ------------------------------------------------------------------ *)
  5.3689 +(*----------------------- norm_Rational_mg ---------------------------*)
  5.3690 +(*. description of the simplifier see MG-DA.p.56ff .*)
  5.3691 +(* ------------------------------------------------------------------- *)
  5.3692 +val common_nominator_p_rls = prep_rls(
  5.3693 +  Rls {id = "common_nominator_p_rls", preconds = [],
  5.3694 +       rew_ord = ("dummy_ord",dummy_ord), 
  5.3695 +	 erls = e_rls, srls = Erls, calc = [],
  5.3696 +	 rules = 
  5.3697 +	 [Rls_ common_nominator_p
  5.3698 +	  (*FIXME.WN0401 ? redesign Rrls - use exhaustively on a term ?
  5.3699 +	    FIXME.WN0510 unnecessary nesting: introduce RRls_ : rls -> rule*)
  5.3700 +	  ], 
  5.3701 +	 scr = EmptyScr});
  5.3702 +(* ------------------------------------------------------------------- *)
  5.3703 +val cancel_p_rls = prep_rls(
  5.3704 +  Rls {id = "cancel_p_rls", preconds = [],
  5.3705 +       rew_ord = ("dummy_ord",dummy_ord), 
  5.3706 +	 erls = e_rls, srls = Erls, calc = [],
  5.3707 +	 rules = 
  5.3708 +	 [Rls_ cancel_p
  5.3709 +	  (*FIXME.WN.0401 ? redesign Rrls - use exhaustively on a term ?*)
  5.3710 +	  ], 
  5.3711 +	 scr = EmptyScr});
  5.3712 +(* -------------------------------------------------------------------- *)
  5.3713 +(*. makes 'normal' fractions; 'is_polyexp' inhibits double fractions;
  5.3714 +    used in initial part norm_Rational_mg, see example DA-M02-main.p.60.*)
  5.3715 +val rat_mult_poly = prep_rls(
  5.3716 +  Rls {id = "rat_mult_poly", preconds = [],
  5.3717 +       rew_ord = ("dummy_ord",dummy_ord), 
  5.3718 +	 erls =  append_rls "e_rls-is_polyexp" e_rls
  5.3719 +	         [Calc ("Poly.is'_polyexp", eval_is_polyexp "")], 
  5.3720 +	 srls = Erls, calc = [],
  5.3721 +	 rules = 
  5.3722 +	 [Thm ("rat_mult_poly_l",num_str rat_mult_poly_l),
  5.3723 +	  (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
  5.3724 +	  Thm ("rat_mult_poly_r",num_str rat_mult_poly_r)
  5.3725 +	  (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*)
  5.3726 +	  ], 
  5.3727 +	 scr = EmptyScr});
  5.3728 +(* ------------------------------------------------------------------ *)
  5.3729 +(*. makes 'normal' fractions; 'is_polyexp' inhibits double fractions;
  5.3730 +    used in looping part norm_Rational_rls, see example DA-M02-main.p.60 
  5.3731 +    .. WHERE THE LATTER DOES ALWAYS WORK, BECAUSE erls = e_rls, 
  5.3732 +    I.E. THE RESPECTIVE ASSUMPTION IS STORED AND Thm APPLIED; WN051028 
  5.3733 +    ... WN0609???MG.*)
  5.3734 +val rat_mult_div_pow = prep_rls(
  5.3735 +  Rls {id = "rat_mult_div_pow", preconds = [], 
  5.3736 +       rew_ord = ("dummy_ord",dummy_ord), 
  5.3737 +       erls = e_rls,
  5.3738 +       (*FIXME.WN051028 append_rls "e_rls-is_polyexp" e_rls
  5.3739 +			[Calc ("Poly.is'_polyexp", eval_is_polyexp "")],
  5.3740 +         with this correction ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ we get 
  5.3741 +	 error "rational.sml.sml: diff.behav. in norm_Rational_mg 29" etc.
  5.3742 +         thus we decided to go on with this flaw*)
  5.3743 +		 srls = Erls, calc = [],
  5.3744 +      rules = [Thm ("rat_mult",num_str rat_mult),
  5.3745 +	       (*"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
  5.3746 +	       Thm ("rat_mult_poly_l",num_str rat_mult_poly_l),
  5.3747 +	       (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
  5.3748 +	       Thm ("rat_mult_poly_r",num_str rat_mult_poly_r),
  5.3749 +	       (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*)
  5.3750 +
  5.3751 +	       Thm ("real_divide_divide1_mg", real_divide_divide1_mg),
  5.3752 +	       (*"y ~= 0 ==> (u / v) / (y / z) = (u * z) / (y * v)"*)
  5.3753 +	       Thm ("real_divide_divide1_eq", real_divide_divide1_eq),
  5.3754 +	       (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
  5.3755 +	       Thm ("real_divide_divide2_eq", real_divide_divide2_eq),
  5.3756 +	       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
  5.3757 +	       Calc ("HOL.divide"  ,eval_cancel "#divide_"),
  5.3758 +	      
  5.3759 +	       Thm ("rat_power", num_str rat_power)
  5.3760 +		(*"(?a / ?b) ^^^ ?n = ?a ^^^ ?n / ?b ^^^ ?n"*)
  5.3761 +	       ],
  5.3762 +      scr = Script ((term_of o the o (parse thy)) "empty_script")
  5.3763 +      }:rls);
  5.3764 +(* ------------------------------------------------------------------ *)
  5.3765 +val rat_reduce_1 = prep_rls(
  5.3766 +  Rls {id = "rat_reduce_1", preconds = [], 
  5.3767 +       rew_ord = ("dummy_ord",dummy_ord), 
  5.3768 +       erls = e_rls, srls = Erls, calc = [], 
  5.3769 +       rules = [Thm ("real_divide_1",num_str real_divide_1),
  5.3770 +		(*"?x / 1 = ?x"*)
  5.3771 +		Thm ("real_mult_1",num_str real_mult_1)           
  5.3772 +		(*"1 * z = z"*)
  5.3773 +		],
  5.3774 +       scr = Script ((term_of o the o (parse thy)) "empty_script")
  5.3775 +       }:rls);
  5.3776 +(* ------------------------------------------------------------------ *)
  5.3777 +(*. looping part of norm_Rational(*_mg*) .*)
  5.3778 +val norm_Rational_rls = prep_rls(
  5.3779 +   Rls {id = "norm_Rational_rls", preconds = [], 
  5.3780 +       rew_ord = ("dummy_ord",dummy_ord), 
  5.3781 +       erls = norm_rat_erls, srls = Erls, calc = [],
  5.3782 +       rules = [Rls_ common_nominator_p_rls,
  5.3783 +		Rls_ rat_mult_div_pow,
  5.3784 +		Rls_ make_rat_poly_with_parentheses,
  5.3785 +		Rls_ cancel_p_rls,(*FIXME:cancel_p does NOT order sometimes*)
  5.3786 +		Rls_ rat_reduce_1
  5.3787 +		],
  5.3788 +       scr = Script ((term_of o the o (parse thy)) "empty_script")
  5.3789 +       }:rls);
  5.3790 +(* ------------------------------------------------------------------ *)
  5.3791 +(*040109 'norm_Rational'(by RL) replaced by 'norm_Rational_mg'(MG)
  5.3792 + just be renaming:*)
  5.3793 +val norm_Rational(*_mg*) = prep_rls(
  5.3794 +   Seq {id = "norm_Rational"(*_mg*), preconds = [], 
  5.3795 +       rew_ord = ("dummy_ord",dummy_ord), 
  5.3796 +       erls = norm_rat_erls, srls = Erls, calc = [],
  5.3797 +       rules = [Rls_ discard_minus_,
  5.3798 +		Rls_ rat_mult_poly,(* removes double fractions like a/b/c    *)
  5.3799 +		Rls_ make_rat_poly_with_parentheses, (*WN0510 also in(#)below*)
  5.3800 +		Rls_ cancel_p_rls, (*FIXME.MG:cancel_p does NOT order sometim*)
  5.3801 +		Rls_ norm_Rational_rls,   (* the main rls, looping (#)       *)
  5.3802 +		Rls_ discard_parentheses_ (* mult only                       *)
  5.3803 +		],
  5.3804 +       scr = Script ((term_of o the o (parse thy)) "empty_script")
  5.3805 +       }:rls);
  5.3806 +(* ------------------------------------------------------------------ *)
  5.3807 +
  5.3808 +
  5.3809 +ruleset' := overwritelthy thy (!ruleset',
  5.3810 +  [("calculate_Rational", calculate_Rational),
  5.3811 +   ("calc_rat_erls",calc_rat_erls),
  5.3812 +   ("rational_erls", rational_erls),
  5.3813 +   ("cancel_p", cancel_p),
  5.3814 +   ("cancel", cancel),
  5.3815 +   ("common_nominator_p", common_nominator_p),
  5.3816 +   ("common_nominator_p_rls", common_nominator_p_rls),
  5.3817 +   ("common_nominator"  , common_nominator),
  5.3818 +   ("discard_minus", discard_minus),
  5.3819 +   ("powers_erls", powers_erls),
  5.3820 +   ("powers", powers),
  5.3821 +   ("rat_mult_divide", rat_mult_divide),
  5.3822 +   ("reduce_0_1_2", reduce_0_1_2),
  5.3823 +   ("rat_reduce_1", rat_reduce_1),
  5.3824 +   ("norm_rat_erls", norm_rat_erls),
  5.3825 +   ("norm_Rational", norm_Rational),
  5.3826 +   ("norm_Rational_rls", norm_Rational_rls),
  5.3827 +   ("norm_Rational_parenthesized", norm_Rational_parenthesized),
  5.3828 +   ("rat_mult_poly", rat_mult_poly),
  5.3829 +   ("rat_mult_div_pow", rat_mult_div_pow),
  5.3830 +   ("cancel_p_rls", cancel_p_rls)
  5.3831 +   ]);
  5.3832 +
  5.3833 +calclist':= overwritel (!calclist', 
  5.3834 +   [("is_expanded", ("Rational.is'_expanded", eval_is_expanded ""))
  5.3835 +    ]);
  5.3836 +
  5.3837 +(** problems **)
  5.3838 +
  5.3839 +store_pbt
  5.3840 + (prep_pbt Rational.thy "pbl_simp_rat" [] e_pblID
  5.3841 + (["rational","simplification"],
  5.3842 +  [("#Given" ,["term t_"]),
  5.3843 +   ("#Where" ,["t_ is_ratpolyexp"]),
  5.3844 +   ("#Find"  ,["normalform n_"])
  5.3845 +  ],
  5.3846 +  append_rls "e_rls" e_rls [(*for preds in where_*)], 
  5.3847 +  SOME "Simplify t_", 
  5.3848 +  [["simplification","of_rationals"]]));
  5.3849 +
  5.3850 +(** methods **)
  5.3851 +
  5.3852 +(*WN061025 this methods script is copied from (auto-generated) script
  5.3853 +  of norm_Rational in order to ease repair on inform*)
  5.3854 +store_met
  5.3855 +    (prep_met Rational.thy "met_simp_rat" [] e_metID
  5.3856 +	      (["simplification","of_rationals"],
  5.3857 +	       [("#Given" ,["term t_"]),
  5.3858 +		("#Where" ,["t_ is_ratpolyexp"]),
  5.3859 +		("#Find"  ,["normalform n_"])
  5.3860 +		],
  5.3861 +	       {rew_ord'="tless_true",
  5.3862 +		rls' = e_rls,
  5.3863 +		calc = [], srls = e_rls, 
  5.3864 +		prls = append_rls "simplification_of_rationals_prls" e_rls 
  5.3865 +				[(*for preds in where_*)
  5.3866 +				 Calc ("Rational.is'_ratpolyexp", 
  5.3867 +				       eval_is_ratpolyexp "")],
  5.3868 +		crls = e_rls, nrls = norm_Rational_rls},
  5.3869 +"Script SimplifyScript (t_::real) =                              " ^
  5.3870 +"  ((Try (Rewrite_Set discard_minus_ False) @@                   " ^
  5.3871 +"    Try (Rewrite_Set rat_mult_poly False) @@                    " ^
  5.3872 +"    Try (Rewrite_Set make_rat_poly_with_parentheses False) @@   " ^
  5.3873 +"    Try (Rewrite_Set cancel_p_rls False) @@                     " ^
  5.3874 +"    (Repeat                                                     " ^
  5.3875 +"     ((Try (Rewrite_Set common_nominator_p_rls False) @@        " ^
  5.3876 +"       Try (Rewrite_Set rat_mult_div_pow False) @@              " ^
  5.3877 +"       Try (Rewrite_Set make_rat_poly_with_parentheses False) @@" ^
  5.3878 +"       Try (Rewrite_Set cancel_p_rls False) @@                  " ^
  5.3879 +"       Try (Rewrite_Set rat_reduce_1 False)))) @@               " ^
  5.3880 +"    Try (Rewrite_Set discard_parentheses_ False))               " ^
  5.3881 +"    t_)"
  5.3882 +	       ));
  5.3883 +*}
  5.3884 +
  5.3885 +end
     6.1 --- a/src/Tools/isac/Knowledge/Root.thy	Thu Aug 26 10:03:53 2010 +0200
     6.2 +++ b/src/Tools/isac/Knowledge/Root.thy	Thu Aug 26 18:15:30 2010 +0200
     6.3 @@ -6,48 +6,329 @@
     6.4               date: 02.10.21
     6.5  *)
     6.6  
     6.7 -(* use_thy_only"Knowledge/Root";
     6.8 -   remove_thy"Root";
     6.9 -   use_thy"Knowledge/Isac";
    6.10 -*)
    6.11 -Root = Simplify + 
    6.12 +theory Root imports Simplify begin
    6.13  
    6.14 -(*-------------------- consts------------------------------------------------*)
    6.15  consts
    6.16  
    6.17 -  sqrt             :: "real => real"         (*"(sqrt _ )" [80] 80*)
    6.18 -  nroot            :: "[real, real] => real"
    6.19 +  sqrt   :: "real => real"         (*"(sqrt _ )" [80] 80*)
    6.20 +  nroot  :: "[real, real] => real"
    6.21  
    6.22 -(*----------------------scripts-----------------------*)
    6.23 -
    6.24 -(*-------------------- rules------------------------------------------------*)
    6.25 -rules (*.not contained in Isabelle2002,
    6.26 +axioms (*.not contained in Isabelle2002,
    6.27           stated as axioms, TODO: prove as theorems;
    6.28           theorem-IDs 'xxxI' with ^^^ instead of ^ in 'xxx' in Isabelle2002.*)
    6.29  
    6.30 -  root_plus_minus       "0 <= b ==> \
    6.31 -			\(a^^^2 = b) = ((a = sqrt b) | (a = (-1)*sqrt b))"
    6.32 -  root_false		"b < 0 ==> (a^^^2 = b) = False"
    6.33 +  root_plus_minus         "0 <= b ==> 
    6.34 +			   (a^^^2 = b) = ((a = sqrt b) | (a = (-1)*sqrt b))"
    6.35 +  root_false		  "b < 0 ==> (a^^^2 = b) = False"
    6.36  
    6.37   (* for expand_rootbinom *)
    6.38 -  real_pp_binom_times        "(a + b)*(c + d) = a*c + a*d + b*c + b*d"
    6.39 -  real_pm_binom_times        "(a + b)*(c - d) = a*c - a*d + b*c - b*d"
    6.40 -  real_mp_binom_times        "(a - b)*(c + d) = a*c + a*d - b*c - b*d"
    6.41 -  real_mm_binom_times        "(a - b)*(c - d) = a*c - a*d - b*c + b*d"
    6.42 -  real_plus_binom_pow3       "(a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3"
    6.43 -  real_minus_binom_pow3      "(a - b)^^^3 = a^^^3 - 3*a^^^2*b + 3*a*b^^^2 - b^^^3"
    6.44 -  realpow_mul                "(a*b)^^^n = a^^^n * b^^^n"
    6.45 +  real_pp_binom_times     "(a + b)*(c + d) = a*c + a*d + b*c + b*d"
    6.46 +  real_pm_binom_times     "(a + b)*(c - d) = a*c - a*d + b*c - b*d"
    6.47 +  real_mp_binom_times     "(a - b)*(c + d) = a*c + a*d - b*c - b*d"
    6.48 +  real_mm_binom_times     "(a - b)*(c - d) = a*c - a*d - b*c + b*d"
    6.49 +  real_plus_binom_pow3    "(a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3"
    6.50 +  real_minus_binom_pow3   "(a - b)^^^3 = a^^^3 - 3*a^^^2*b + 3*a*b^^^2 - b^^^3"
    6.51 +  realpow_mul             "(a*b)^^^n = a^^^n * b^^^n"
    6.52  
    6.53 -  real_diff_minus            "a - b = a + (-1) * b"
    6.54 -  real_plus_binom_times      "(a + b)*(a + b) = a^^^2 + 2*a*b + b^^^2"
    6.55 -  real_minus_binom_times     "(a - b)*(a - b) = a^^^2 - 2*a*b + b^^^2"
    6.56 -  real_plus_binom_pow2       "(a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"
    6.57 -  real_minus_binom_pow2      "(a - b)^^^2 = a^^^2 - 2*a*b + b^^^2"
    6.58 -  real_plus_minus_binom1     "(a + b)*(a - b) = a^^^2 - b^^^2"
    6.59 -  real_plus_minus_binom2     "(a - b)*(a + b) = a^^^2 - b^^^2"
    6.60 +  real_diff_minus         "a - b = a + (-1) * b"
    6.61 +  real_plus_binom_times   "(a + b)*(a + b) = a^^^2 + 2*a*b + b^^^2"
    6.62 +  real_minus_binom_times  "(a - b)*(a - b) = a^^^2 - 2*a*b + b^^^2"
    6.63 +  real_plus_binom_pow2    "(a + b)^^^2 = a^^^2 + 2*a*b + b^^^2"
    6.64 +  real_minus_binom_pow2   "(a - b)^^^2 = a^^^2 - 2*a*b + b^^^2"
    6.65 +  real_plus_minus_binom1  "(a + b)*(a - b) = a^^^2 - b^^^2"
    6.66 +  real_plus_minus_binom2  "(a - b)*(a + b) = a^^^2 - b^^^2"
    6.67  
    6.68 -  real_root_positive     "0 <= a ==> (x ^^^ 2 = a) = (x = sqrt a)"
    6.69 -  real_root_negative     "a <  0 ==> (x ^^^ 2 = a) = False"
    6.70 +  real_root_positive      "0 <= a ==> (x ^^^ 2 = a) = (x = sqrt a)"
    6.71 +  real_root_negative      "a <  0 ==> (x ^^^ 2 = a) = False"
    6.72  
    6.73 -      
    6.74 +ML {*
    6.75 +(*-------------------------functions---------------------*)
    6.76 +(*evaluation square-root over the integers*)
    6.77 +fun eval_sqrt (thmid:string) (op_:string) (t as 
    6.78 +	       (Const(op0,t0) $ arg)) thy = 
    6.79 +    (case arg of 
    6.80 +	Free (n1,t1) =>
    6.81 +	(case int_of_str n1 of
    6.82 +	     SOME ni => 
    6.83 +	     if ni < 0 then NONE
    6.84 +	     else
    6.85 +		 let val fact = squfact ni;
    6.86 +		 in if fact*fact = ni 
    6.87 +		    then SOME ("#sqrt #"^(string_of_int ni)^" = #"
    6.88 +			       ^(string_of_int (if ni = 0 then 0
    6.89 +						else ni div fact)),
    6.90 +			       Trueprop $ mk_equality (t, term_of_num t1 fact))
    6.91 +		    else if fact = 1 then NONE
    6.92 +		    else SOME ("#sqrt #"^(string_of_int ni)^" = sqrt (#"
    6.93 +			       ^(string_of_int fact)^" * #"
    6.94 +			       ^(string_of_int fact)^" * #"
    6.95 +			       ^(string_of_int (ni div (fact*fact))^")"),
    6.96 +			       Trueprop $ 
    6.97 +					(mk_equality 
    6.98 +					     (t, 
    6.99 +					      (mk_factroot op0 t1 fact 
   6.100 +						(ni div (fact*fact))))))
   6.101 +	         end
   6.102 +	   | NONE => NONE)
   6.103 +      | _ => NONE)
   6.104 +
   6.105 +  | eval_sqrt _ _ _ _ = NONE;
   6.106 +(*val (thmid, op_, t as Const(op0,t0) $ arg) = ("","", str2term "sqrt 0");
   6.107 +> eval_sqrt thmid op_ t thy;
   6.108 +> val Free (n1,t1) = arg; 
   6.109 +> val SOME ni = int_of_str n1;
   6.110 +*)
   6.111 +
   6.112 +calclist':= overwritel (!calclist', 
   6.113 +   [("SQRT"    ,("Root.sqrt"   ,eval_sqrt "#sqrt_"))
   6.114 +    (*different types for 'sqrt 4' --- 'Calculate sqrt_'*)
   6.115 +    ]);
   6.116 +
   6.117 +
   6.118 +local (* Vers. 7.10.99.A *)
   6.119 +
   6.120 +open Term;  (* for type order = EQUAL | LESS | GREATER *)
   6.121 +
   6.122 +fun pr_ord EQUAL = "EQUAL"
   6.123 +  | pr_ord LESS  = "LESS"
   6.124 +  | pr_ord GREATER = "GREATER";
   6.125 +
   6.126 +fun dest_hd' (Const (a, T)) =                          (* ~ term.ML *)
   6.127 +  (case a of "Root.sqrt"  => ((("|||", 0), T), 0)      (*WN greatest *)
   6.128 +	   | _ => (((a, 0), T), 0))
   6.129 +  | dest_hd' (Free (a, T)) = (((a, 0), T), 1)
   6.130 +  | dest_hd' (Var v) = (v, 2)
   6.131 +  | dest_hd' (Bound i) = ((("", i), dummyT), 3)
   6.132 +  | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4);
   6.133 +fun size_of_term' (Const(str,_) $ t) =
   6.134 +    (case str of "Root.sqrt"  => (1000 + size_of_term' t)
   6.135 +               | _ => 1 + size_of_term' t)
   6.136 +  | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
   6.137 +  | size_of_term' (f $ t) = size_of_term' f  +  size_of_term' t
   6.138 +  | size_of_term' _ = 1;
   6.139 +fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
   6.140 +      (case term_ord' pr thy (t, u) of EQUAL => typ_ord (T, U) | ord => ord)
   6.141 +  | term_ord' pr thy (t, u) =
   6.142 +      (if pr then 
   6.143 +	 let
   6.144 +	   val (f, ts) = strip_comb t and (g, us) = strip_comb u;
   6.145 +	   val _=writeln("t= f@ts= \""^
   6.146 +	      ((Syntax.string_of_term (thy2ctxt thy)) f)^"\" @ \"["^
   6.147 +	      (commas(map(Syntax.string_of_term (thy2ctxt thy)) ts))^"]\"");
   6.148 +	   val _=writeln("u= g@us= \""^
   6.149 +	      ((Syntax.string_of_term (thy2ctxt thy)) g)^"\" @ \"["^
   6.150 +	      (commas(map(Syntax.string_of_term (thy2ctxt thy)) us))^"]\"");
   6.151 +	   val _=writeln("size_of_term(t,u)= ("^
   6.152 +	      (string_of_int(size_of_term' t))^", "^
   6.153 +	      (string_of_int(size_of_term' u))^")");
   6.154 +	   val _=writeln("hd_ord(f,g)      = "^((pr_ord o hd_ord)(f,g)));
   6.155 +	   val _=writeln("terms_ord(ts,us) = "^
   6.156 +			   ((pr_ord o terms_ord str false)(ts,us)));
   6.157 +	   val _=writeln("-------");
   6.158 +	 in () end
   6.159 +       else ();
   6.160 +	 case int_ord (size_of_term' t, size_of_term' u) of
   6.161 +	   EQUAL =>
   6.162 +	     let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
   6.163 +	       (case hd_ord (f, g) of EQUAL => (terms_ord str pr) (ts, us) 
   6.164 +	     | ord => ord)
   6.165 +	     end
   6.166 +	 | ord => ord)
   6.167 +and hd_ord (f, g) =                                        (* ~ term.ML *)
   6.168 +  prod_ord (prod_ord indexname_ord typ_ord) int_ord (dest_hd' f, dest_hd' g)
   6.169 +and terms_ord str pr (ts, us) = 
   6.170 +    list_ord (term_ord' pr (assoc_thy "Isac.thy"))(ts, us);
   6.171 +
   6.172 +in
   6.173 +(* associates a+(b+c) => (a+b)+c = a+b+c ... avoiding parentheses 
   6.174 +  by (1) size_of_term: less(!) to right, size_of 'sqrt (...)' = 1 
   6.175 +     (2) hd_ord: greater to right, 'sqrt' < numerals < variables
   6.176 +     (3) terms_ord: recurs. on args, greater to right
   6.177 +*)
   6.178 +
   6.179 +(*args
   6.180 +   pr: print trace, WN0509 'sqrt_right true' not used anymore
   6.181 +   thy:
   6.182 +   subst: no bound variables, only Root.sqrt
   6.183 +   tu: the terms to compare (t1, t2) ... *)
   6.184 +fun sqrt_right (pr:bool) thy (_:subst) tu = 
   6.185 +    (term_ord' pr thy(***) tu = LESS );
   6.186 +end;
   6.187 +
   6.188 +rew_ord' := overwritel (!rew_ord',
   6.189 +[("termlessI", termlessI),
   6.190 + ("sqrt_right", sqrt_right false (theory "Pure"))
   6.191 + ]);
   6.192 +
   6.193 +(*-------------------------rulse-------------------------*)
   6.194 +val Root_crls = 
   6.195 +      append_rls "Root_crls" Atools_erls 
   6.196 +       [Thm  ("real_unari_minus",num_str real_unari_minus),
   6.197 +        Calc ("Root.sqrt" ,eval_sqrt "#sqrt_"),
   6.198 +        Calc ("HOL.divide",eval_cancel "#divide_"),
   6.199 +        Calc ("Atools.pow" ,eval_binop "#power_"),
   6.200 +        Calc ("op +", eval_binop "#add_"), 
   6.201 +        Calc ("op -", eval_binop "#sub_"),
   6.202 +        Calc ("op *", eval_binop "#mult_"),
   6.203 +        Calc ("op =",eval_equal "#equal_") 
   6.204 +        ];
   6.205 +
   6.206 +val Root_erls = 
   6.207 +      append_rls "Root_erls" Atools_erls 
   6.208 +       [Thm  ("real_unari_minus",num_str real_unari_minus),
   6.209 +        Calc ("Root.sqrt" ,eval_sqrt "#sqrt_"),
   6.210 +        Calc ("HOL.divide",eval_cancel "#divide_"),
   6.211 +        Calc ("Atools.pow" ,eval_binop "#power_"),
   6.212 +        Calc ("op +", eval_binop "#add_"), 
   6.213 +        Calc ("op -", eval_binop "#sub_"),
   6.214 +        Calc ("op *", eval_binop "#mult_"),
   6.215 +        Calc ("op =",eval_equal "#equal_") 
   6.216 +        ];
   6.217 +
   6.218 +ruleset' := overwritelthy thy (!ruleset',
   6.219 +			[("Root_erls",Root_erls) (*FIXXXME:del with rls.rls'*) 
   6.220 +			 ]);
   6.221 +
   6.222 +val make_rooteq = prep_rls(
   6.223 +  Rls{id = "make_rooteq", preconds = []:term list, 
   6.224 +      rew_ord = ("sqrt_right", sqrt_right false Root.thy),
   6.225 +      erls = Atools_erls, srls = Erls,
   6.226 +      calc = [],
   6.227 +      (*asm_thm = [],*)
   6.228 +      rules = [Thm ("real_diff_minus",num_str real_diff_minus),			
   6.229 +	       (*"a - b = a + (-1) * b"*)
   6.230 +
   6.231 +	       Thm ("real_add_mult_distrib" ,num_str real_add_mult_distrib),	
   6.232 +	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
   6.233 +	       Thm ("real_add_mult_distrib2",num_str real_add_mult_distrib2),	
   6.234 +	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
   6.235 +	       Thm ("real_diff_mult_distrib" ,num_str real_diff_mult_distrib),	
   6.236 +	       (*"(z1.0 - z2.0) * w = z1.0 * w - z2.0 * w"*)
   6.237 +	       Thm ("real_diff_mult_distrib2",num_str real_diff_mult_distrib2),	
   6.238 +	       (*"w * (z1.0 - z2.0) = w * z1.0 - w * z2.0"*)
   6.239 +
   6.240 +	       Thm ("real_mult_1",num_str real_mult_1),                         
   6.241 +	       (*"1 * z = z"*)
   6.242 +	       Thm ("real_mult_0",num_str real_mult_0),                         
   6.243 +	       (*"0 * z = 0"*)
   6.244 +	       Thm ("real_add_zero_left",num_str real_add_zero_left),		
   6.245 +	       (*"0 + z = z"*)
   6.246 + 
   6.247 +	       Thm ("real_mult_commute",num_str real_mult_commute),
   6.248 +		(*AC-rewriting*)
   6.249 +	       Thm ("real_mult_left_commute",num_str real_mult_left_commute),
   6.250 +         	(**)
   6.251 +	       Thm ("real_mult_assoc",num_str real_mult_assoc),
   6.252 +	        (**)
   6.253 +	       Thm ("real_add_commute",num_str real_add_commute),
   6.254 +		(**)
   6.255 +	       Thm ("real_add_left_commute",num_str real_add_left_commute),
   6.256 +	        (**)
   6.257 +	       Thm ("real_add_assoc",num_str real_add_assoc),
   6.258 +	        (**)
   6.259 +
   6.260 +	       Thm ("sym_realpow_twoI",num_str (realpow_twoI RS sym)),		
   6.261 +	       (*"r1 * r1 = r1 ^^^ 2"*)
   6.262 +	       Thm ("realpow_plus_1",num_str realpow_plus_1),			
   6.263 +	       (*"r * r ^^^ n = r ^^^ (n + 1)"*)
   6.264 +	       Thm ("sym_real_mult_2",num_str (real_mult_2 RS sym)),		
   6.265 +	       (*"z1 + z1 = 2 * z1"*)
   6.266 +	       Thm ("real_mult_2_assoc",num_str real_mult_2_assoc),		
   6.267 +	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
   6.268 +
   6.269 +	       Thm ("real_num_collect",num_str real_num_collect), 
   6.270 +	       (*"[| l is_const; m is_const |]==> l * n + m * n = (l + m) * n"*)
   6.271 +	       Thm ("real_num_collect_assoc",num_str real_num_collect_assoc),	
   6.272 +	       (*"[| l is_const; m is_const |] ==>  
   6.273 +                                   l * n + (m * n + k) =  (l + m) * n + k"*)
   6.274 +	       Thm ("real_one_collect",num_str real_one_collect),		
   6.275 +	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
   6.276 +	       Thm ("real_one_collect_assoc",num_str real_one_collect_assoc), 
   6.277 +	       (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
   6.278 +
   6.279 +	       Calc ("op +", eval_binop "#add_"), 
   6.280 +	       Calc ("op *", eval_binop "#mult_"),
   6.281 +	       Calc ("Atools.pow", eval_binop "#power_")
   6.282 +	       ],
   6.283 +      scr = Script ((term_of o the o (parse thy)) "empty_script")
   6.284 +      }:rls);      
   6.285 +ruleset' := overwritelthy thy (!ruleset',
   6.286 +			[("make_rooteq", make_rooteq)
   6.287 +			 ]);
   6.288 +
   6.289 +val expand_rootbinoms = prep_rls(
   6.290 +  Rls{id = "expand_rootbinoms", preconds = [], 
   6.291 +      rew_ord = ("termlessI",termlessI),
   6.292 +      erls = Atools_erls, srls = Erls,
   6.293 +      calc = [],
   6.294 +      (*asm_thm = [],*)
   6.295 +      rules = [Thm ("real_plus_binom_pow2"  ,num_str real_plus_binom_pow2),     
   6.296 +	       (*"(a + b) ^^^ 2 = a ^^^ 2 + 2 * a * b + b ^^^ 2"*)
   6.297 +	       Thm ("real_plus_binom_times" ,num_str real_plus_binom_times),    
   6.298 +	       (*"(a + b)*(a + b) = ...*)
   6.299 +	       Thm ("real_minus_binom_pow2" ,num_str real_minus_binom_pow2),    
   6.300 +		(*"(a - b) ^^^ 2 = a ^^^ 2 - 2 * a * b + b ^^^ 2"*)
   6.301 +	       Thm ("real_minus_binom_times",num_str real_minus_binom_times),   
   6.302 +	       (*"(a - b)*(a - b) = ...*)
   6.303 +	       Thm ("real_plus_minus_binom1",num_str real_plus_minus_binom1),   
   6.304 +		(*"(a + b) * (a - b) = a ^^^ 2 - b ^^^ 2"*)
   6.305 +	       Thm ("real_plus_minus_binom2",num_str real_plus_minus_binom2),   
   6.306 +		(*"(a - b) * (a + b) = a ^^^ 2 - b ^^^ 2"*)
   6.307 +	       (*RL 020915*)
   6.308 +	       Thm ("real_pp_binom_times",num_str real_pp_binom_times), 
   6.309 +		(*(a + b)*(c + d) = a*c + a*d + b*c + b*d*)
   6.310 +               Thm ("real_pm_binom_times",num_str real_pm_binom_times), 
   6.311 +		(*(a + b)*(c - d) = a*c - a*d + b*c - b*d*)
   6.312 +               Thm ("real_mp_binom_times",num_str real_mp_binom_times), 
   6.313 +		(*(a - b)*(c p d) = a*c + a*d - b*c - b*d*)
   6.314 +               Thm ("real_mm_binom_times",num_str real_mm_binom_times), 
   6.315 +		(*(a - b)*(c p d) = a*c - a*d - b*c + b*d*)
   6.316 +	       Thm ("realpow_mul",num_str realpow_mul),                 
   6.317 +		(*(a*b)^^^n = a^^^n * b^^^n*)
   6.318 +
   6.319 +	       Thm ("real_mult_1",num_str real_mult_1),         (*"1 * z = z"*)
   6.320 +	       Thm ("real_mult_0",num_str real_mult_0),         (*"0 * z = 0"*)
   6.321 +	       Thm ("real_add_zero_left",num_str real_add_zero_left), 
   6.322 +                 (*"0 + z = z"*)
   6.323 +
   6.324 +	       Calc ("op +", eval_binop "#add_"), 
   6.325 +	       Calc ("op -", eval_binop "#sub_"), 
   6.326 +	       Calc ("op *", eval_binop "#mult_"),
   6.327 +	       Calc ("HOL.divide"  ,eval_cancel "#divide_"),
   6.328 +	       Calc ("Root.sqrt",eval_sqrt "#sqrt_"),
   6.329 +	       Calc ("Atools.pow", eval_binop "#power_"),
   6.330 +
   6.331 +	       Thm ("sym_realpow_twoI",num_str (realpow_twoI RS sym)),		
   6.332 +	       (*"r1 * r1 = r1 ^^^ 2"*)
   6.333 +	       Thm ("realpow_plus_1",num_str realpow_plus_1),			
   6.334 +	       (*"r * r ^^^ n = r ^^^ (n + 1)"*)
   6.335 +	       Thm ("real_mult_2_assoc",num_str real_mult_2_assoc),		
   6.336 +	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
   6.337 +
   6.338 +	       Thm ("real_num_collect",num_str real_num_collect), 
   6.339 +	       (*"[| l is_const; m is_const |] ==>l * n + m * n = (l + m) * n"*)
   6.340 +	       Thm ("real_num_collect_assoc",num_str real_num_collect_assoc),	
   6.341 +	       (*"[| l is_const; m is_const |] ==>
   6.342 +                  l * n + (m * n + k) =  (l + m) * n + k"*)
   6.343 +	       Thm ("real_one_collect",num_str real_one_collect),		
   6.344 +	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
   6.345 +	       Thm ("real_one_collect_assoc",num_str real_one_collect_assoc), 
   6.346 +	       (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
   6.347 +
   6.348 +	       Calc ("op +", eval_binop "#add_"), 
   6.349 +	       Calc ("op -", eval_binop "#sub_"), 
   6.350 +	       Calc ("op *", eval_binop "#mult_"),
   6.351 +	       Calc ("HOL.divide"  ,eval_cancel "#divide_"),
   6.352 +	       Calc ("Root.sqrt",eval_sqrt "#sqrt_"),
   6.353 +	       Calc ("Atools.pow", eval_binop "#power_")
   6.354 +	       ],
   6.355 +      scr = Script ((term_of o the o (parse thy)) "empty_script")
   6.356 +       }:rls);      
   6.357 +
   6.358 +
   6.359 +ruleset' := overwritelthy thy (!ruleset',
   6.360 +			[("expand_rootbinoms", expand_rootbinoms)
   6.361 +			 ]);
   6.362 +*}
   6.363 +
   6.364  end
     7.1 --- a/src/Tools/isac/Knowledge/RootEq.thy	Thu Aug 26 10:03:53 2010 +0200
     7.2 +++ b/src/Tools/isac/Knowledge/RootEq.thy	Thu Aug 26 18:15:30 2010 +0200
     7.3 @@ -6,137 +6,645 @@
     7.4     last change by: rlang
     7.5               date: 02.11.14
     7.6  *)
     7.7 -(*  use"../knowledge/RootEq.ML";
     7.8 -   use"knowledge/RootEq.ML";
     7.9 -   use"RootEq.ML";
    7.10  
    7.11 -   remove_thy"RootEq";
    7.12 -   use_thy"Isac";
    7.13 +theory RootEq imports Root end
    7.14  
    7.15 -   use"ROOT.ML";
    7.16 -   cd"knowledge";
    7.17 - *)
    7.18 -
    7.19 -RootEq = Root + 
    7.20 -
    7.21 -(*-------------------- consts------------------------------------------------*)
    7.22  consts
    7.23    (*-------------------------root-----------------------*)
    7.24 -  is'_rootTerm'_in :: [real, real] => bool ("_ is'_rootTerm'_in _") 
    7.25 -  is'_sqrtTerm'_in :: [real, real] => bool ("_ is'_sqrtTerm'_in _") 
    7.26 -  is'_normSqrtTerm'_in :: [real, real] => bool ("_ is'_normSqrtTerm'_in _") 
    7.27 +  is'_rootTerm'_in     :: "[real, real] => bool" ("_ is'_rootTerm'_in _") 
    7.28 +  is'_sqrtTerm'_in     :: "[real, real] => bool" ("_ is'_sqrtTerm'_in _") 
    7.29 +  is'_normSqrtTerm'_in :: "[real, real] => bool" ("_ is'_normSqrtTerm'_in _") 
    7.30 +
    7.31    (*----------------------scripts-----------------------*)
    7.32    Norm'_sq'_root'_equation
    7.33 -             :: "[bool,real, \
    7.34 -		  \ bool list] => bool list"
    7.35 -               ("((Script Norm'_sq'_root'_equation (_ _ =))// \
    7.36 -                 \ (_))" 9)
    7.37 +             :: "[bool,real, 
    7.38 +		   bool list] => bool list"
    7.39 +               ("((Script Norm'_sq'_root'_equation (_ _ =))// 
    7.40 +                  (_))" 9)
    7.41    Solve'_sq'_root'_equation
    7.42 -             :: "[bool,real, \
    7.43 -		  \ bool list] => bool list"
    7.44 -               ("((Script Solve'_sq'_root'_equation (_ _ =))// \
    7.45 -                 \ (_))" 9)
    7.46 +             :: "[bool,real, 
    7.47 +		   bool list] => bool list"
    7.48 +               ("((Script Solve'_sq'_root'_equation (_ _ =))// 
    7.49 +                  (_))" 9)
    7.50    Solve'_left'_sq'_root'_equation
    7.51 -             :: "[bool,real, \
    7.52 -		  \ bool list] => bool list"
    7.53 -               ("((Script Solve'_left'_sq'_root'_equation (_ _ =))// \
    7.54 -                 \ (_))" 9)
    7.55 +             :: "[bool,real, 
    7.56 +		   bool list] => bool list"
    7.57 +               ("((Script Solve'_left'_sq'_root'_equation (_ _ =))// 
    7.58 +                  (_))" 9)
    7.59    Solve'_right'_sq'_root'_equation
    7.60 -             :: "[bool,real, \
    7.61 -		  \ bool list] => bool list"
    7.62 -               ("((Script Solve'_right'_sq'_root'_equation (_ _ =))// \
    7.63 -                 \ (_))" 9)
    7.64 +             :: "[bool,real, 
    7.65 +		   bool list] => bool list"
    7.66 +               ("((Script Solve'_right'_sq'_root'_equation (_ _ =))// 
    7.67 +                  (_))" 9)
    7.68   
    7.69 -(*-------------------- rules------------------------------------------------*)
    7.70 -rules 
    7.71 +axioms 
    7.72  
    7.73  (* normalize *)
    7.74 -  makex1_x
    7.75 -    "a^^^1  = a"  
    7.76 -  real_assoc_1
    7.77 -   "a+(b+c) = a+b+c"
    7.78 -  real_assoc_2
    7.79 -   "a*(b*c) = a*b*c"
    7.80 +  makex1_x            "a^^^1  = a"  
    7.81 +  real_assoc_1        "a+(b+c) = a+b+c"
    7.82 +  real_assoc_2        "a*(b*c) = a*b*c"
    7.83  
    7.84    (* simplification of root*)
    7.85 -  sqrt_square_1
    7.86 -  "[|0 <= a|] ==>  (sqrt a)^^^2 = a"
    7.87 -  sqrt_square_2
    7.88 -   "sqrt (a ^^^ 2) = a"
    7.89 -  sqrt_times_root_1
    7.90 -   "sqrt a * sqrt b = sqrt(a*b)"
    7.91 -  sqrt_times_root_2
    7.92 -   "a * sqrt b * sqrt c = a * sqrt(b*c)"
    7.93 +  sqrt_square_1       "[|0 <= a|] ==>  (sqrt a)^^^2 = a"
    7.94 +  sqrt_square_2       "sqrt (a ^^^ 2) = a"
    7.95 +  sqrt_times_root_1   "sqrt a * sqrt b = sqrt(a*b)"
    7.96 +  sqrt_times_root_2   "a * sqrt b * sqrt c = a * sqrt(b*c)"
    7.97  
    7.98    (* isolate one root on the LEFT or RIGHT hand side of the equation *)
    7.99 -  sqrt_isolate_l_add1
   7.100 -  "[|bdv occurs_in c|] ==> (a + b*sqrt(c) = d) = (b * sqrt(c) = d+ (-1) * a)"
   7.101 -  sqrt_isolate_l_add2
   7.102 -  "[|bdv occurs_in c|] ==>(a + sqrt(c) = d) = ((sqrt(c) = d+ (-1) * a))"
   7.103 -  sqrt_isolate_l_add3
   7.104 -  "[|bdv occurs_in c|] ==> (a + b*(e/sqrt(c)) = d) = (b * (e/sqrt(c)) = d+ (-1) * a)"
   7.105 -  sqrt_isolate_l_add4
   7.106 -  "[|bdv occurs_in c|] ==>(a + b/(f*sqrt(c)) = d) = (b / (f*sqrt(c)) = d+ (-1) * a)"
   7.107 -  sqrt_isolate_l_add5
   7.108 -  "[|bdv occurs_in c|] ==> (a + b*(e/(f*sqrt(c))) = d) = (b * (e/(f*sqrt(c))) = d+ (-1) * a)"
   7.109 -  sqrt_isolate_l_add6
   7.110 -  "[|bdv occurs_in c|] ==>(a + b/sqrt(c) = d) = (b / sqrt(c) = d+ (-1) * a)"
   7.111 -  sqrt_isolate_r_add1
   7.112 -  "[|bdv occurs_in f|] ==>(a = d + e*sqrt(f)) = (a + (-1) * d = e*sqrt(f))"
   7.113 -  sqrt_isolate_r_add2
   7.114 -  "[|bdv occurs_in f|] ==>(a = d + sqrt(f)) = (a + (-1) * d = sqrt(f))"
   7.115 +  sqrt_isolate_l_add1 "[|bdv occurs_in c|] ==> 
   7.116 +   (a + b*sqrt(c) = d) = (b * sqrt(c) = d+ (-1) * a)"
   7.117 +  sqrt_isolate_l_add2 "[|bdv occurs_in c|] ==>
   7.118 +   (a + sqrt(c) = d) = ((sqrt(c) = d+ (-1) * a))"
   7.119 +  sqrt_isolate_l_add3 "[|bdv occurs_in c|] ==>
   7.120 +   (a + b*(e/sqrt(c)) = d) = (b * (e/sqrt(c)) = d + (-1) * a)"
   7.121 +  sqrt_isolate_l_add4 "[|bdv occurs_in c|] ==>
   7.122 +   (a + b/(f*sqrt(c)) = d) = (b / (f*sqrt(c)) = d + (-1) * a)"
   7.123 +  sqrt_isolate_l_add5 "[|bdv occurs_in c|] ==>
   7.124 +   (a + b*(e/(f*sqrt(c))) = d) = (b * (e/(f*sqrt(c))) = d+ (-1) * a)"
   7.125 +  sqrt_isolate_l_add6 "[|bdv occurs_in c|] ==>
   7.126 +   (a + b/sqrt(c) = d) = (b / sqrt(c) = d+ (-1) * a)"
   7.127 +  sqrt_isolate_r_add1 "[|bdv occurs_in f|] ==>
   7.128 +   (a = d + e*sqrt(f)) = (a + (-1) * d = e*sqrt(f))"
   7.129 +  sqrt_isolate_r_add2 "[|bdv occurs_in f|] ==>
   7.130 +   (a = d + sqrt(f)) = (a + (-1) * d = sqrt(f))"
   7.131   (* small hack: thm 3,5,6 are not needed if rootnormalize is well done*)
   7.132 -  sqrt_isolate_r_add3
   7.133 -  "[|bdv occurs_in f|] ==>(a = d + e*(g/sqrt(f))) = (a + (-1) * d = e*(g/sqrt(f)))"
   7.134 -  sqrt_isolate_r_add4
   7.135 -  "[|bdv occurs_in f|] ==>(a = d + g/sqrt(f)) = (a + (-1) * d = g/sqrt(f))"
   7.136 -  sqrt_isolate_r_add5
   7.137 -  "[|bdv occurs_in f|] ==>(a = d + e*(g/(h*sqrt(f)))) = (a + (-1) * d = e*(g/(h*sqrt(f))))"
   7.138 -  sqrt_isolate_r_add6
   7.139 -  "[|bdv occurs_in f|] ==>(a = d + g/(h*sqrt(f))) = (a + (-1) * d = g/(h*sqrt(f)))"
   7.140 +  sqrt_isolate_r_add3 "[|bdv occurs_in f|] ==>
   7.141 +   (a = d + e*(g/sqrt(f))) = (a + (-1) * d = e*(g/sqrt(f)))"
   7.142 +  sqrt_isolate_r_add4 "[|bdv occurs_in f|] ==>
   7.143 +   (a = d + g/sqrt(f)) = (a + (-1) * d = g/sqrt(f))"
   7.144 +  sqrt_isolate_r_add5 "[|bdv occurs_in f|] ==>
   7.145 +   (a = d + e*(g/(h*sqrt(f)))) = (a + (-1) * d = e*(g/(h*sqrt(f))))"
   7.146 +  sqrt_isolate_r_add6 "[|bdv occurs_in f|] ==>
   7.147 +   (a = d + g/(h*sqrt(f))) = (a + (-1) * d = g/(h*sqrt(f)))"
   7.148   
   7.149    (* eliminate isolates sqrt *)
   7.150 -  sqrt_square_equation_both_1
   7.151 -  "[|bdv occurs_in b; bdv occurs_in d|] ==> 
   7.152 -               ( (sqrt a + sqrt b         = sqrt c + sqrt d) = 
   7.153 -                 (a+2*sqrt(a)*sqrt(b)+b  = c+2*sqrt(c)*sqrt(d)+d))"
   7.154 -  sqrt_square_equation_both_2
   7.155 -  "[|bdv occurs_in b; bdv occurs_in d|] ==> 
   7.156 -               ( (sqrt a - sqrt b           = sqrt c + sqrt d) = 
   7.157 -                 (a - 2*sqrt(a)*sqrt(b)+b  = c+2*sqrt(c)*sqrt(d)+d))"
   7.158 -  sqrt_square_equation_both_3
   7.159 -  "[|bdv occurs_in b; bdv occurs_in d|] ==> 
   7.160 -               ( (sqrt a + sqrt b           = sqrt c - sqrt d) = 
   7.161 -                 (a + 2*sqrt(a)*sqrt(b)+b  = c - 2*sqrt(c)*sqrt(d)+d))"
   7.162 -  sqrt_square_equation_both_4
   7.163 -  "[|bdv occurs_in b; bdv occurs_in d|] ==> 
   7.164 -               ( (sqrt a - sqrt b           = sqrt c - sqrt d) = 
   7.165 -                 (a - 2*sqrt(a)*sqrt(b)+b  = c - 2*sqrt(c)*sqrt(d)+d))"
   7.166 -  sqrt_square_equation_left_1
   7.167 -  "[|bdv occurs_in a; 0 <= a; 0 <= b|] ==> ( (sqrt (a) = b) = (a = (b^^^2)))"
   7.168 -  sqrt_square_equation_left_2
   7.169 -  "[|bdv occurs_in a; 0 <= a; 0 <= b*c|] ==> ( (c*sqrt(a) = b) = (c^^^2*a = b^^^2))"
   7.170 -  sqrt_square_equation_left_3
   7.171 -  "[|bdv occurs_in a; 0 <= a; 0 <= b*c|] ==> ( c/sqrt(a) = b) = (c^^^2 / a = b^^^2)"
   7.172 +  sqrt_square_equation_both_1 "[|bdv occurs_in b; bdv occurs_in d|] ==> 
   7.173 +   ( (sqrt a + sqrt b         = sqrt c + sqrt d) = 
   7.174 +     (a+2*sqrt(a)*sqrt(b)+b  = c+2*sqrt(c)*sqrt(d)+d))"
   7.175 +  sqrt_square_equation_both_2 "[|bdv occurs_in b; bdv occurs_in d|] ==> 
   7.176 +   ( (sqrt a - sqrt b           = sqrt c + sqrt d) = 
   7.177 +     (a - 2*sqrt(a)*sqrt(b)+b  = c+2*sqrt(c)*sqrt(d)+d))"
   7.178 +  sqrt_square_equation_both_3 "[|bdv occurs_in b; bdv occurs_in d|] ==> 
   7.179 +   ( (sqrt a + sqrt b           = sqrt c - sqrt d) = 
   7.180 +     (a + 2*sqrt(a)*sqrt(b)+b  = c - 2*sqrt(c)*sqrt(d)+d))"
   7.181 +  sqrt_square_equation_both_4 "[|bdv occurs_in b; bdv occurs_in d|] ==> 
   7.182 +   ( (sqrt a - sqrt b           = sqrt c - sqrt d) = 
   7.183 +     (a - 2*sqrt(a)*sqrt(b)+b  = c - 2*sqrt(c)*sqrt(d)+d))"
   7.184 +  sqrt_square_equation_left_1 "[|bdv occurs_in a; 0 <= a; 0 <= b|] ==>
   7.185 +   ( (sqrt (a) = b) = (a = (b^^^2)))"
   7.186 +  sqrt_square_equation_left_2 "[|bdv occurs_in a; 0 <= a; 0 <= b*c|] ==> 
   7.187 +   ( (c*sqrt(a) = b) = (c^^^2*a = b^^^2))"
   7.188 +  sqrt_square_equation_left_3 "[|bdv occurs_in a; 0 <= a; 0 <= b*c|] ==> 
   7.189 +   ( c/sqrt(a) = b) = (c^^^2 / a = b^^^2)"
   7.190    (* small hack: thm 4-6 are not needed if rootnormalize is well done*)
   7.191 -  sqrt_square_equation_left_4
   7.192 -  "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d|] ==> ( (c*(d/sqrt (a)) = b) = (c^^^2*(d^^^2/a) = b^^^2))"
   7.193 -  sqrt_square_equation_left_5
   7.194 -  "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d|] ==> ( c/(d*sqrt(a)) = b) = (c^^^2 / (d^^^2*a) = b^^^2)"
   7.195 -  sqrt_square_equation_left_6
   7.196 -  "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d*e|] ==> ( (c*(d/(e*sqrt (a))) = b) = (c^^^2*(d^^^2/(e^^^2*a)) = b^^^2))"
   7.197 -  sqrt_square_equation_right_1
   7.198 -  "[|bdv occurs_in b; 0 <= a; 0 <= b|] ==> ( (a = sqrt (b)) = (a^^^2 = b))"
   7.199 -  sqrt_square_equation_right_2
   7.200 -  "[|bdv occurs_in b; 0 <= a*c; 0 <= b|] ==> ( (a = c*sqrt (b)) = ((a^^^2) = c^^^2*b))"
   7.201 -  sqrt_square_equation_right_3
   7.202 -  "[|bdv occurs_in b; 0 <= a*c; 0 <= b|] ==> ( (a = c/sqrt (b)) = (a^^^2 = c^^^2/b))"
   7.203 +  sqrt_square_equation_left_4 "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d|] ==> 
   7.204 +   ( (c*(d/sqrt (a)) = b) = (c^^^2*(d^^^2/a) = b^^^2))"
   7.205 +  sqrt_square_equation_left_5 "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d|] ==> 
   7.206 +   ( c/(d*sqrt(a)) = b) = (c^^^2 / (d^^^2*a) = b^^^2)"
   7.207 +  sqrt_square_equation_left_6 "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d*e|] ==> 
   7.208 +   ( (c*(d/(e*sqrt (a))) = b) = (c^^^2*(d^^^2/(e^^^2*a)) = b^^^2))"
   7.209 +  sqrt_square_equation_right_1  "[|bdv occurs_in b; 0 <= a; 0 <= b|] ==> 
   7.210 +   ( (a = sqrt (b)) = (a^^^2 = b))"
   7.211 +  sqrt_square_equation_right_2 "[|bdv occurs_in b; 0 <= a*c; 0 <= b|] ==> 
   7.212 +   ( (a = c*sqrt (b)) = ((a^^^2) = c^^^2*b))"
   7.213 +  sqrt_square_equation_right_3 "[|bdv occurs_in b; 0 <= a*c; 0 <= b|] ==> 
   7.214 +   ( (a = c/sqrt (b)) = (a^^^2 = c^^^2/b))"
   7.215   (* small hack: thm 4-6 are not needed if rootnormalize is well done*)
   7.216 -  sqrt_square_equation_right_4
   7.217 -  "[|bdv occurs_in b; 0 <= a*c*d; 0 <= b|] ==> ( (a = c*(d/sqrt (b))) = ((a^^^2) = c^^^2*(d^^^2/b)))"
   7.218 -  sqrt_square_equation_right_5
   7.219 -  "[|bdv occurs_in b; 0 <= a*c*d; 0 <= b|] ==> ( (a = c/(d*sqrt (b))) = (a^^^2 = c^^^2/(d^^^2*b)))"
   7.220 -  sqrt_square_equation_right_6
   7.221 -  "[|bdv occurs_in b; 0 <= a*c*d*e; 0 <= b|] ==> ( (a = c*(d/(e*sqrt (b)))) = ((a^^^2) = c^^^2*(d^^^2/(e^^^2*b))))"
   7.222 - 
   7.223 +  sqrt_square_equation_right_4 "[|bdv occurs_in b; 0 <= a*c*d; 0 <= b|] ==> 
   7.224 +   ( (a = c*(d/sqrt (b))) = ((a^^^2) = c^^^2*(d^^^2/b)))"
   7.225 +  sqrt_square_equation_right_5 "[|bdv occurs_in b; 0 <= a*c*d; 0 <= b|] ==> 
   7.226 +   ( (a = c/(d*sqrt (b))) = (a^^^2 = c^^^2/(d^^^2*b)))"
   7.227 +  sqrt_square_equation_right_6 "[|bdv occurs_in b; 0 <= a*c*d*e; 0 <= b|] ==> 
   7.228 +   ( (a = c*(d/(e*sqrt (b)))) = ((a^^^2) = c^^^2*(d^^^2/(e^^^2*b))))"
   7.229 +
   7.230 +ML {*
   7.231 +(*-------------------------functions---------------------*)
   7.232 +(* true if bdv is under sqrt of a Equation*)
   7.233 +fun is_rootTerm_in t v = 
   7.234 +    let 
   7.235 +	fun coeff_in c v = member op = (vars c) v;
   7.236 +   	fun findroot (_ $ _ $ _ $ _) v = raise error("is_rootTerm_in:")
   7.237 +	  (* at the moment there is no term like this, but ....*)
   7.238 +	  | findroot (t as (Const ("Root.nroot",_) $ _ $ t3)) v = coeff_in t3 v
   7.239 +	  | findroot (_ $ t2 $ t3) v = (findroot t2 v) orelse (findroot t3 v)
   7.240 +	  | findroot (t as (Const ("Root.sqrt",_) $ t2)) v = coeff_in t2 v
   7.241 +	  | findroot (_ $ t2) v = (findroot t2 v)
   7.242 +	  | findroot _ _ = false;
   7.243 +     in
   7.244 +	findroot t v
   7.245 +    end;
   7.246 +
   7.247 + fun is_sqrtTerm_in t v = 
   7.248 +    let 
   7.249 +	fun coeff_in c v = member op = (vars c) v;
   7.250 +   	fun findsqrt (_ $ _ $ _ $ _) v = raise error("is_sqrteqation_in:")
   7.251 +	  (* at the moment there is no term like this, but ....*)
   7.252 +	  | findsqrt (_ $ t1 $ t2) v = (findsqrt t1 v) orelse (findsqrt t2 v)
   7.253 +	  | findsqrt (t as (Const ("Root.sqrt",_) $ a)) v = coeff_in a v
   7.254 +	  | findsqrt (_ $ t1) v = (findsqrt t1 v)
   7.255 +	  | findsqrt _ _ = false;
   7.256 +     in
   7.257 +	findsqrt t v
   7.258 +    end;
   7.259 +
   7.260 +(* RL: 030518: Is in the rightest subterm of a term a sqrt with bdv,
   7.261 +and the subterm ist connected with + or * --> is normalized*)
   7.262 + fun is_normSqrtTerm_in t v =
   7.263 +     let
   7.264 +	fun coeff_in c v = member op = (vars c) v;
   7.265 +        fun isnorm (_ $ _ $ _ $ _) v = raise error("is_normSqrtTerm_in:")
   7.266 +	  (* at the moment there is no term like this, but ....*)
   7.267 +          | isnorm (Const ("op +",_) $ _ $ t2) v = is_sqrtTerm_in t2 v
   7.268 +          | isnorm (Const ("op *",_) $ _ $ t2) v = is_sqrtTerm_in t2 v
   7.269 +          | isnorm (Const ("op -",_) $ _ $ _) v = false
   7.270 +          | isnorm (Const ("HOL.divide",_) $ t1 $ t2) v = (is_sqrtTerm_in t1 v) orelse 
   7.271 +                              (is_sqrtTerm_in t2 v)
   7.272 +          | isnorm (Const ("Root.sqrt",_) $ t1) v = coeff_in t1 v
   7.273 + 	  | isnorm (_ $ t1) v = is_sqrtTerm_in t1 v
   7.274 +          | isnorm _ _ = false;
   7.275 +     in
   7.276 +         isnorm t v
   7.277 +     end;
   7.278 +
   7.279 +fun eval_is_rootTerm_in _ _ 
   7.280 +       (p as (Const ("RootEq.is'_rootTerm'_in",_) $ t $ v)) _  =
   7.281 +    if is_rootTerm_in t v then 
   7.282 +	SOME ((term2str p) ^ " = True",
   7.283 +	      Trueprop $ (mk_equality (p, HOLogic.true_const)))
   7.284 +    else SOME ((term2str p) ^ " = True",
   7.285 +	       Trueprop $ (mk_equality (p, HOLogic.false_const)))
   7.286 +  | eval_is_rootTerm_in _ _ _ _ = ((*writeln"### nichts matcht";*) NONE);
   7.287 +
   7.288 +fun eval_is_sqrtTerm_in _ _ 
   7.289 +       (p as (Const ("RootEq.is'_sqrtTerm'_in",_) $ t $ v)) _  =
   7.290 +    if is_sqrtTerm_in t v then 
   7.291 +	SOME ((term2str p) ^ " = True",
   7.292 +	      Trueprop $ (mk_equality (p, HOLogic.true_const)))
   7.293 +    else SOME ((term2str p) ^ " = True",
   7.294 +	       Trueprop $ (mk_equality (p, HOLogic.false_const)))
   7.295 +  | eval_is_sqrtTerm_in _ _ _ _ = ((*writeln"### nichts matcht";*) NONE);
   7.296 +
   7.297 +fun eval_is_normSqrtTerm_in _ _ 
   7.298 +       (p as (Const ("RootEq.is'_normSqrtTerm'_in",_) $ t $ v)) _  =
   7.299 +    if is_normSqrtTerm_in t v then 
   7.300 +	SOME ((term2str p) ^ " = True",
   7.301 +	      Trueprop $ (mk_equality (p, HOLogic.true_const)))
   7.302 +    else SOME ((term2str p) ^ " = True",
   7.303 +	       Trueprop $ (mk_equality (p, HOLogic.false_const)))
   7.304 +  | eval_is_normSqrtTerm_in _ _ _ _ = ((*writeln"### nichts matcht";*) NONE);
   7.305 +
   7.306 +(*-------------------------rulse-------------------------*)
   7.307 +val RootEq_prls =(*15.10.02:just the following order due to subterm evaluation*)
   7.308 +  append_rls "RootEq_prls" e_rls 
   7.309 +	     [Calc ("Atools.ident",eval_ident "#ident_"),
   7.310 +	      Calc ("Tools.matches",eval_matches ""),
   7.311 +	      Calc ("Tools.lhs"    ,eval_lhs ""),
   7.312 +	      Calc ("Tools.rhs"    ,eval_rhs ""),
   7.313 +	      Calc ("RootEq.is'_sqrtTerm'_in",eval_is_sqrtTerm_in ""),
   7.314 +	      Calc ("RootEq.is'_rootTerm'_in",eval_is_rootTerm_in ""),
   7.315 +	      Calc ("RootEq.is'_normSqrtTerm'_in",eval_is_normSqrtTerm_in ""),
   7.316 +	      Calc ("op =",eval_equal "#equal_"),
   7.317 +	      Thm ("not_true",num_str not_true),
   7.318 +	      Thm ("not_false",num_str not_false),
   7.319 +	      Thm ("and_true",num_str and_true),
   7.320 +	      Thm ("and_false",num_str and_false),
   7.321 +	      Thm ("or_true",num_str or_true),
   7.322 +	      Thm ("or_false",num_str or_false)
   7.323 +	      ];
   7.324 +
   7.325 +val RootEq_erls =
   7.326 +     append_rls "RootEq_erls" Root_erls
   7.327 +          [Thm ("real_divide_divide2_eq",num_str real_divide_divide2_eq)
   7.328 +           ];
   7.329 +
   7.330 +val RootEq_crls = 
   7.331 +     append_rls "RootEq_crls" Root_crls
   7.332 +          [Thm ("real_divide_divide2_eq",num_str real_divide_divide2_eq)
   7.333 +           ];
   7.334 +
   7.335 +val rooteq_srls = 
   7.336 +     append_rls "rooteq_srls" e_rls
   7.337 +		[Calc ("RootEq.is'_sqrtTerm'_in",eval_is_sqrtTerm_in ""),
   7.338 +                 Calc ("RootEq.is'_normSqrtTerm'_in",eval_is_normSqrtTerm_in""),
   7.339 +                 Calc ("RootEq.is'_rootTerm'_in",eval_is_rootTerm_in "")
   7.340 +		 ];
   7.341 +
   7.342 +ruleset' := overwritelthy thy (!ruleset',
   7.343 +			[("RootEq_erls",RootEq_erls),
   7.344 +                                             (*FIXXXME:del with rls.rls'*)
   7.345 +			 ("rooteq_srls",rooteq_srls)
   7.346 +                         ]);
   7.347 +
   7.348 +(*isolate the bound variable in an sqrt equation; 'bdv' is a meta-constant*)
   7.349 + val sqrt_isolate = prep_rls(
   7.350 +  Rls {id = "sqrt_isolate", preconds = [], rew_ord = ("termlessI",termlessI), 
   7.351 +       erls = RootEq_erls, srls = Erls, calc = [], 
   7.352 +       rules = [
   7.353 +       Thm("sqrt_square_1",num_str sqrt_square_1),
   7.354 +                     (* (sqrt a)^^^2 -> a *)
   7.355 +       Thm("sqrt_square_2",num_str sqrt_square_2),
   7.356 +                     (* sqrt (a^^^2) -> a *)
   7.357 +       Thm("sqrt_times_root_1",num_str sqrt_times_root_1),
   7.358 +            (* sqrt a sqrt b -> sqrt(ab) *)
   7.359 +       Thm("sqrt_times_root_2",num_str sqrt_times_root_2),
   7.360 +            (* a sqrt b sqrt c -> a sqrt(bc) *)
   7.361 +       Thm("sqrt_square_equation_both_1",
   7.362 +           num_str sqrt_square_equation_both_1),
   7.363 +       (* (sqrt a + sqrt b  = sqrt c + sqrt d) -> 
   7.364 +            (a+2*sqrt(a)*sqrt(b)+b) = c+2*sqrt(c)*sqrt(d)+d) *)
   7.365 +       Thm("sqrt_square_equation_both_2",
   7.366 +            num_str sqrt_square_equation_both_2),
   7.367 +       (* (sqrt a - sqrt b  = sqrt c + sqrt d) -> 
   7.368 +            (a-2*sqrt(a)*sqrt(b)+b) = c+2*sqrt(c)*sqrt(d)+d) *)
   7.369 +       Thm("sqrt_square_equation_both_3",
   7.370 +            num_str sqrt_square_equation_both_3),
   7.371 +       (* (sqrt a + sqrt b  = sqrt c - sqrt d) -> 
   7.372 +            (a+2*sqrt(a)*sqrt(b)+b) = c-2*sqrt(c)*sqrt(d)+d) *)
   7.373 +       Thm("sqrt_square_equation_both_4",
   7.374 +            num_str sqrt_square_equation_both_4),
   7.375 +       (* (sqrt a - sqrt b  = sqrt c - sqrt d) -> 
   7.376 +            (a-2*sqrt(a)*sqrt(b)+b) = c-2*sqrt(c)*sqrt(d)+d) *)
   7.377 +       Thm("sqrt_isolate_l_add1",
   7.378 +            num_str sqrt_isolate_l_add1), 
   7.379 +       (* a+b*sqrt(x)=d -> b*sqrt(x) = d-a *)
   7.380 +       Thm("sqrt_isolate_l_add2",
   7.381 +            num_str sqrt_isolate_l_add2), 
   7.382 +       (* a+  sqrt(x)=d ->   sqrt(x) = d-a *)
   7.383 +       Thm("sqrt_isolate_l_add3",
   7.384 +            num_str sqrt_isolate_l_add3), 
   7.385 +       (* a+b*c/sqrt(x)=d->b*c/sqrt(x)=d-a *)
   7.386 +       Thm("sqrt_isolate_l_add4",
   7.387 +            num_str sqrt_isolate_l_add4), 
   7.388 +       (* a+c/sqrt(x)=d -> c/sqrt(x) = d-a *)
   7.389 +       Thm("sqrt_isolate_l_add5",
   7.390 +            num_str sqrt_isolate_l_add5), 
   7.391 +       (* a+b*c/f*sqrt(x)=d->b*c/f*sqrt(x)=d-a *)
   7.392 +       Thm("sqrt_isolate_l_add6",
   7.393 +            num_str sqrt_isolate_l_add6), 
   7.394 +       (* a+c/f*sqrt(x)=d -> c/f*sqrt(x) = d-a *)
   7.395 +       (*Thm("sqrt_isolate_l_div",num_str sqrt_isolate_l_div),*)
   7.396 +         (* b*sqrt(x) = d sqrt(x) d/b *)
   7.397 +       Thm("sqrt_isolate_r_add1",
   7.398 +            num_str sqrt_isolate_r_add1),
   7.399 +       (* a= d+e*sqrt(x) -> a-d=e*sqrt(x) *)
   7.400 +       Thm("sqrt_isolate_r_add2",
   7.401 +            num_str sqrt_isolate_r_add2),
   7.402 +       (* a= d+  sqrt(x) -> a-d=  sqrt(x) *)
   7.403 +       Thm("sqrt_isolate_r_add3",
   7.404 +            num_str sqrt_isolate_r_add3),
   7.405 +       (* a=d+e*g/sqrt(x)->a-d=e*g/sqrt(x)*)
   7.406 +       Thm("sqrt_isolate_r_add4",
   7.407 +            num_str sqrt_isolate_r_add4),
   7.408 +       (* a= d+g/sqrt(x) -> a-d=g/sqrt(x) *)
   7.409 +       Thm("sqrt_isolate_r_add5",
   7.410 +            num_str sqrt_isolate_r_add5),
   7.411 +       (* a=d+e*g/h*sqrt(x)->a-d=e*g/h*sqrt(x)*)
   7.412 +       Thm("sqrt_isolate_r_add6",
   7.413 +            num_str sqrt_isolate_r_add6),
   7.414 +       (* a= d+g/h*sqrt(x) -> a-d=g/h*sqrt(x) *)
   7.415 +       (*Thm("sqrt_isolate_r_div",num_str sqrt_isolate_r_div),*)
   7.416 +         (* a=e*sqrt(x) -> a/e = sqrt(x) *)
   7.417 +       Thm("sqrt_square_equation_left_1",
   7.418 +            num_str sqrt_square_equation_left_1),   
   7.419 +       (* sqrt(x)=b -> x=b^2 *)
   7.420 +       Thm("sqrt_square_equation_left_2",
   7.421 +            num_str sqrt_square_equation_left_2),   
   7.422 +       (* c*sqrt(x)=b -> c^2*x=b^2 *)
   7.423 +       Thm("sqrt_square_equation_left_3",num_str sqrt_square_equation_left_3),  
   7.424 +	      (* c/sqrt(x)=b -> c^2/x=b^2 *)
   7.425 +       Thm("sqrt_square_equation_left_4",num_str sqrt_square_equation_left_4),
   7.426 +	      (* c*d/sqrt(x)=b -> c^2*d^2/x=b^2 *)
   7.427 +       Thm("sqrt_square_equation_left_5",num_str sqrt_square_equation_left_5),
   7.428 +	      (* c/d*sqrt(x)=b -> c^2/d^2x=b^2 *)
   7.429 +       Thm("sqrt_square_equation_left_6",num_str sqrt_square_equation_left_6),
   7.430 +	      (* c*d/g*sqrt(x)=b -> c^2*d^2/g^2x=b^2 *)
   7.431 +       Thm("sqrt_square_equation_right_1",num_str sqrt_square_equation_right_1),
   7.432 +	      (* a=sqrt(x) ->a^2=x *)
   7.433 +       Thm("sqrt_square_equation_right_2",num_str sqrt_square_equation_right_2),
   7.434 +	      (* a=c*sqrt(x) ->a^2=c^2*x *)
   7.435 +       Thm("sqrt_square_equation_right_3",num_str sqrt_square_equation_right_3),
   7.436 +	      (* a=c/sqrt(x) ->a^2=c^2/x *)
   7.437 +       Thm("sqrt_square_equation_right_4",num_str sqrt_square_equation_right_4),
   7.438 +	      (* a=c*d/sqrt(x) ->a^2=c^2*d^2/x *)
   7.439 +       Thm("sqrt_square_equation_right_5",num_str sqrt_square_equation_right_5),
   7.440 +	      (* a=c/e*sqrt(x) ->a^2=c^2/e^2x *)
   7.441 +       Thm("sqrt_square_equation_right_6",num_str sqrt_square_equation_right_6)
   7.442 +	      (* a=c*d/g*sqrt(x) ->a^2=c^2*d^2/g^2*x *)
   7.443 +       ],scr = Script ((term_of o the o (parse thy)) "empty_script")
   7.444 +      }:rls);
   7.445 +
   7.446 +ruleset' := overwritelthy thy (!ruleset',
   7.447 +			[("sqrt_isolate",sqrt_isolate)
   7.448 +			 ]);
   7.449 +(* -- left 28.08.02--*)
   7.450 +(*isolate the bound variable in an sqrt left equation; 'bdv' is a meta-constant*)
   7.451 + val l_sqrt_isolate = prep_rls(
   7.452 +     Rls {id = "l_sqrt_isolate", preconds = [], 
   7.453 +	  rew_ord = ("termlessI",termlessI), 
   7.454 +          erls = RootEq_erls, srls = Erls, calc = [], 
   7.455 +     rules = [
   7.456 +     Thm("sqrt_square_1",num_str sqrt_square_1),
   7.457 +                            (* (sqrt a)^^^2 -> a *)
   7.458 +     Thm("sqrt_square_2",num_str sqrt_square_2),
   7.459 +                            (* sqrt (a^^^2) -> a *)
   7.460 +     Thm("sqrt_times_root_1",num_str sqrt_times_root_1),
   7.461 +            (* sqrt a sqrt b -> sqrt(ab) *)
   7.462 +     Thm("sqrt_times_root_2",num_str sqrt_times_root_2),
   7.463 +        (* a sqrt b sqrt c -> a sqrt(bc) *)
   7.464 +     Thm("sqrt_isolate_l_add1",num_str sqrt_isolate_l_add1),
   7.465 +        (* a+b*sqrt(x)=d -> b*sqrt(x) = d-a *)
   7.466 +     Thm("sqrt_isolate_l_add2",num_str sqrt_isolate_l_add2),
   7.467 +        (* a+  sqrt(x)=d ->   sqrt(x) = d-a *)
   7.468 +     Thm("sqrt_isolate_l_add3",num_str sqrt_isolate_l_add3),
   7.469 +        (* a+b*c/sqrt(x)=d->b*c/sqrt(x)=d-a *)
   7.470 +     Thm("sqrt_isolate_l_add4",num_str sqrt_isolate_l_add4),
   7.471 +        (* a+c/sqrt(x)=d -> c/sqrt(x) = d-a *)
   7.472 +     Thm("sqrt_isolate_l_add5",num_str sqrt_isolate_l_add5),
   7.473 +        (* a+b*c/f*sqrt(x)=d->b*c/f*sqrt(x)=d-a *)
   7.474 +     Thm("sqrt_isolate_l_add6",num_str sqrt_isolate_l_add6),
   7.475 +        (* a+c/f*sqrt(x)=d -> c/f*sqrt(x) = d-a *)
   7.476 +   (*Thm("sqrt_isolate_l_div",num_str sqrt_isolate_l_div),*)
   7.477 +        (* b*sqrt(x) = d sqrt(x) d/b *)
   7.478 +     Thm("sqrt_square_equation_left_1",num_str sqrt_square_equation_left_1),
   7.479 +	      (* sqrt(x)=b -> x=b^2 *)
   7.480 +     Thm("sqrt_square_equation_left_2",num_str sqrt_square_equation_left_2),
   7.481 +	      (* a*sqrt(x)=b -> a^2*x=b^2*)
   7.482 +     Thm("sqrt_square_equation_left_3",num_str sqrt_square_equation_left_3),   
   7.483 +	      (* c/sqrt(x)=b -> c^2/x=b^2 *)
   7.484 +     Thm("sqrt_square_equation_left_4",num_str sqrt_square_equation_left_4),   
   7.485 +	      (* c*d/sqrt(x)=b -> c^2*d^2/x=b^2 *)
   7.486 +     Thm("sqrt_square_equation_left_5",num_str sqrt_square_equation_left_5),   
   7.487 +	      (* c/d*sqrt(x)=b -> c^2/d^2x=b^2 *)
   7.488 +     Thm("sqrt_square_equation_left_6",num_str sqrt_square_equation_left_6)  
   7.489 +	      (* c*d/g*sqrt(x)=b -> c^2*d^2/g^2x=b^2 *)
   7.490 +    ],
   7.491 +    scr = Script ((term_of o the o (parse thy)) "empty_script")
   7.492 +   }:rls);
   7.493 +
   7.494 +ruleset' := overwritelthy thy (!ruleset',
   7.495 +			[("l_sqrt_isolate",l_sqrt_isolate)
   7.496 +			 ]);
   7.497 +
   7.498 +(* -- right 28.8.02--*)
   7.499 +(*isolate the bound variable in an sqrt right equation; 'bdv' is a meta-constant*)
   7.500 + val r_sqrt_isolate = prep_rls(
   7.501 +     Rls {id = "r_sqrt_isolate", preconds = [], 
   7.502 +	  rew_ord = ("termlessI",termlessI), 
   7.503 +          erls = RootEq_erls, srls = Erls, calc = [], 
   7.504 +     rules = [
   7.505 +     Thm("sqrt_square_1",num_str sqrt_square_1),
   7.506 +                           (* (sqrt a)^^^2 -> a *)
   7.507 +     Thm("sqrt_square_2",num_str sqrt_square_2), 
   7.508 +                           (* sqrt (a^^^2) -> a *)
   7.509 +     Thm("sqrt_times_root_1",num_str sqrt_times_root_1),
   7.510 +           (* sqrt a sqrt b -> sqrt(ab) *)
   7.511 +     Thm("sqrt_times_root_2",num_str sqrt_times_root_2),
   7.512 +       (* a sqrt b sqrt c -> a sqrt(bc) *)
   7.513 +     Thm("sqrt_isolate_r_add1",num_str sqrt_isolate_r_add1),
   7.514 +       (* a= d+e*sqrt(x) -> a-d=e*sqrt(x) *)
   7.515 +     Thm("sqrt_isolate_r_add2",num_str sqrt_isolate_r_add2),
   7.516 +       (* a= d+  sqrt(x) -> a-d=  sqrt(x) *)
   7.517 +     Thm("sqrt_isolate_r_add3",num_str sqrt_isolate_r_add3),
   7.518 +       (* a=d+e*g/sqrt(x)->a-d=e*g/sqrt(x)*)
   7.519 +     Thm("sqrt_isolate_r_add4",num_str sqrt_isolate_r_add4),
   7.520 +       (* a= d+g/sqrt(x) -> a-d=g/sqrt(x) *)
   7.521 +     Thm("sqrt_isolate_r_add5",num_str sqrt_isolate_r_add5),
   7.522 +       (* a=d+e*g/h*sqrt(x)->a-d=e*g/h*sqrt(x)*)
   7.523 +     Thm("sqrt_isolate_r_add6",num_str sqrt_isolate_r_add6),
   7.524 +       (* a= d+g/h*sqrt(x) -> a-d=g/h*sqrt(x) *)
   7.525 +   (*Thm("sqrt_isolate_r_div",num_str sqrt_isolate_r_div),*)
   7.526 +       (* a=e*sqrt(x) -> a/e = sqrt(x) *)
   7.527 +     Thm("sqrt_square_equation_right_1",num_str sqrt_square_equation_right_1),
   7.528 +	      (* a=sqrt(x) ->a^2=x *)
   7.529 +     Thm("sqrt_square_equation_right_2",num_str sqrt_square_equation_right_2),
   7.530 +	      (* a=c*sqrt(x) ->a^2=c^2*x *)
   7.531 +     Thm("sqrt_square_equation_right_3",num_str sqrt_square_equation_right_3),
   7.532 +	      (* a=c/sqrt(x) ->a^2=c^2/x *)
   7.533 +     Thm("sqrt_square_equation_right_4",num_str sqrt_square_equation_right_4), 
   7.534 +	      (* a=c*d/sqrt(x) ->a^2=c^2*d^2/x *)
   7.535 +     Thm("sqrt_square_equation_right_5",num_str sqrt_square_equation_right_5),
   7.536 +	      (* a=c/e*sqrt(x) ->a^2=c^2/e^2x *)
   7.537 +     Thm("sqrt_square_equation_right_6",num_str sqrt_square_equation_right_6)
   7.538 +	      (* a=c*d/g*sqrt(x) ->a^2=c^2*d^2/g^2*x *)
   7.539 +    ],
   7.540 +    scr = Script ((term_of o the o (parse thy)) "empty_script")
   7.541 +   }:rls);
   7.542 +
   7.543 +ruleset' := overwritelthy thy (!ruleset',
   7.544 +			[("r_sqrt_isolate",r_sqrt_isolate)
   7.545 +			 ]);
   7.546 +
   7.547 +val rooteq_simplify = prep_rls(
   7.548 +  Rls {id = "rooteq_simplify", 
   7.549 +       preconds = [], rew_ord = ("termlessI",termlessI), 
   7.550 +       erls = RootEq_erls, srls = Erls, calc = [], 
   7.551 +       (*asm_thm = [("sqrt_square_1","")],*)
   7.552 +       rules = [Thm  ("real_assoc_1",num_str real_assoc_1),
   7.553 +                             (* a+(b+c) = a+b+c *)
   7.554 +                Thm  ("real_assoc_2",num_str real_assoc_2),
   7.555 +                             (* a*(b*c) = a*b*c *)
   7.556 +                Calc ("op +",eval_binop "#add_"),
   7.557 +                Calc ("op -",eval_binop "#sub_"),
   7.558 +                Calc ("op *",eval_binop "#mult_"),
   7.559 +                Calc ("HOL.divide", eval_cancel "#divide_"),
   7.560 +                Calc ("Root.sqrt",eval_sqrt "#sqrt_"),
   7.561 +                Calc ("Atools.pow" ,eval_binop "#power_"),
   7.562 +                Thm("real_plus_binom_pow2",num_str real_plus_binom_pow2),
   7.563 +                Thm("real_minus_binom_pow2",num_str real_minus_binom_pow2),
   7.564 +                Thm("realpow_mul",num_str realpow_mul),    
   7.565 +                     (* (a * b)^n = a^n * b^n*)
   7.566 +                Thm("sqrt_times_root_1",num_str sqrt_times_root_1), 
   7.567 +                     (* sqrt b * sqrt c = sqrt(b*c) *)
   7.568 +                Thm("sqrt_times_root_2",num_str sqrt_times_root_2),
   7.569 +                     (* a * sqrt a * sqrt b = a * sqrt(a*b) *)
   7.570 +                Thm("sqrt_square_2",num_str sqrt_square_2),
   7.571 +                            (* sqrt (a^^^2) = a *)
   7.572 +                Thm("sqrt_square_1",num_str sqrt_square_1) 
   7.573 +                            (* sqrt a ^^^ 2 = a *)
   7.574 +                ],
   7.575 +       scr = Script ((term_of o the o (parse thy)) "empty_script")
   7.576 +    }:rls);
   7.577 +  ruleset' := overwritelthy thy (!ruleset',
   7.578 +                          [("rooteq_simplify",rooteq_simplify)
   7.579 +                           ]);
   7.580 +  
   7.581 +(*-------------------------Problem-----------------------*)
   7.582 +(*
   7.583 +(get_pbt ["root","univariate","equation"]);
   7.584 +show_ptyps(); 
   7.585 +*)
   7.586 +(* ---------root----------- *)
   7.587 +store_pbt
   7.588 + (prep_pbt RootEq.thy "pbl_equ_univ_root" [] e_pblID
   7.589 + (["root","univariate","equation"],
   7.590 +  [("#Given" ,["equality e_","solveFor v_"]),
   7.591 +   ("#Where" ,["(lhs e_) is_rootTerm_in  (v_::real) | " ^
   7.592 +	       "(rhs e_) is_rootTerm_in  (v_::real)"]),
   7.593 +   ("#Find"  ,["solutions v_i_"]) 
   7.594 +  ],
   7.595 +  RootEq_prls, SOME "solve (e_::bool, v_)",
   7.596 +  []));
   7.597 +(* ---------sqrt----------- *)
   7.598 +store_pbt
   7.599 + (prep_pbt RootEq.thy "pbl_equ_univ_root_sq" [] e_pblID
   7.600 + (["sq","root","univariate","equation"],
   7.601 +  [("#Given" ,["equality e_","solveFor v_"]),
   7.602 +   ("#Where" ,["( ((lhs e_) is_sqrtTerm_in (v_::real)) &" ^
   7.603 +               "  ((lhs e_) is_normSqrtTerm_in (v_::real))   )  |" ^
   7.604 +	       "( ((rhs e_) is_sqrtTerm_in (v_::real)) &" ^
   7.605 +               "  ((rhs e_) is_normSqrtTerm_in (v_::real))   )"]),
   7.606 +   ("#Find"  ,["solutions v_i_"]) 
   7.607 +  ],
   7.608 +  RootEq_prls,  SOME "solve (e_::bool, v_)",
   7.609 +  [["RootEq","solve_sq_root_equation"]]));
   7.610 +(* ---------normalize----------- *)
   7.611 +store_pbt
   7.612 + (prep_pbt RootEq.thy "pbl_equ_univ_root_norm" [] e_pblID
   7.613 + (["normalize","root","univariate","equation"],
   7.614 +  [("#Given" ,["equality e_","solveFor v_"]),
   7.615 +   ("#Where" ,["( ((lhs e_) is_sqrtTerm_in (v_::real)) &" ^
   7.616 +               "  Not((lhs e_) is_normSqrtTerm_in (v_::real)))  | " ^
   7.617 +	       "( ((rhs e_) is_sqrtTerm_in (v_::real)) &" ^
   7.618 +               "  Not((rhs e_) is_normSqrtTerm_in (v_::real)))"]),
   7.619 +   ("#Find"  ,["solutions v_i_"]) 
   7.620 +  ],
   7.621 +  RootEq_prls,  SOME "solve (e_::bool, v_)",
   7.622 +  [["RootEq","norm_sq_root_equation"]]));
   7.623 +
   7.624 +(*-------------------------methods-----------------------*)
   7.625 +(* ---- root 20.8.02 ---*)
   7.626 +store_met
   7.627 + (prep_met RootEq.thy "met_rooteq" [] e_metID
   7.628 + (["RootEq"],
   7.629 +   [],
   7.630 +   {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = e_rls, prls=e_rls,
   7.631 +    crls=RootEq_crls, nrls=norm_Poly(*,
   7.632 +    asm_rls=[],asm_thm=[]*)}, "empty_script"));
   7.633 +(*-- normalize 20.10.02 --*)
   7.634 +store_met
   7.635 + (prep_met RootEq.thy "met_rooteq_norm" [] e_metID
   7.636 + (["RootEq","norm_sq_root_equation"],
   7.637 +   [("#Given" ,["equality e_","solveFor v_"]),
   7.638 +    ("#Where" ,["( ((lhs e_) is_sqrtTerm_in (v_::real)) &" ^
   7.639 +               "  Not((lhs e_) is_normSqrtTerm_in (v_::real)))  | " ^
   7.640 +	       "( ((rhs e_) is_sqrtTerm_in (v_::real)) &" ^
   7.641 +               "  Not((rhs e_) is_normSqrtTerm_in (v_::real)))"]),
   7.642 +    ("#Find"  ,["solutions v_i_"])
   7.643 +   ],
   7.644 +   {rew_ord'="termlessI",
   7.645 +    rls'=RootEq_erls,
   7.646 +    srls=e_rls,
   7.647 +    prls=RootEq_prls,
   7.648 +    calc=[],
   7.649 +    crls=RootEq_crls, nrls=norm_Poly(*,
   7.650 +    asm_rls=[],
   7.651 +    asm_thm=[("sqrt_square_1","")]*)},
   7.652 +   "Script Norm_sq_root_equation  (e_::bool) (v_::real)  =                " ^
   7.653 +    "(let e_ = ((Repeat(Try (Rewrite     makex1_x            False))) @@  " ^
   7.654 +    "           (Try (Repeat (Rewrite_Set expand_rootbinoms  False))) @@  " ^ 
   7.655 +    "           (Try (Rewrite_Set rooteq_simplify              True)) @@  " ^ 
   7.656 +    "           (Try (Repeat (Rewrite_Set make_rooteq        False))) @@  " ^
   7.657 +    "           (Try (Rewrite_Set rooteq_simplify              True))) e_ " ^
   7.658 +    " in ((SubProblem (RootEq_,[univariate,equation],                     " ^
   7.659 +    "      [no_met]) [bool_ e_, real_ v_])))"
   7.660 +   ));
   7.661 +
   7.662 +store_met
   7.663 + (prep_met RootEq.thy "met_rooteq_sq" [] e_metID
   7.664 + (["RootEq","solve_sq_root_equation"],
   7.665 +   [("#Given" ,["equality e_","solveFor v_"]),
   7.666 +    ("#Where" ,["( ((lhs e_) is_sqrtTerm_in (v_::real)) &" ^
   7.667 +                "  ((lhs e_) is_normSqrtTerm_in (v_::real))   )  |" ^
   7.668 +	        "( ((rhs e_) is_sqrtTerm_in (v_::real)) &" ^
   7.669 +                "  ((rhs e_) is_normSqrtTerm_in (v_::real))   )"]),
   7.670 +    ("#Find"  ,["solutions v_i_"])
   7.671 +   ],
   7.672 +   {rew_ord'="termlessI",
   7.673 +    rls'=RootEq_erls,
   7.674 +    srls = rooteq_srls,
   7.675 +    prls = RootEq_prls,
   7.676 +    calc = [],
   7.677 +    crls=RootEq_crls, nrls=norm_Poly},
   7.678 +"Script Solve_sq_root_equation  (e_::bool) (v_::real)  =             " ^
   7.679 +"(let e_ = " ^
   7.680 +"  ((Try (Rewrite_Set_Inst [(bdv,v_::real)] sqrt_isolate    True)) @@ " ^
   7.681 +"  (Try (Rewrite_Set                       rooteq_simplify True)) @@ " ^
   7.682 +"  (Try (Repeat (Rewrite_Set expand_rootbinoms           False))) @@ " ^
   7.683 +"  (Try (Repeat (Rewrite_Set make_rooteq                 False))) @@ " ^
   7.684 +"  (Try (Rewrite_Set rooteq_simplify                       True))) e_;" ^
   7.685 +" (L_::bool list) =                                                   " ^
   7.686 +"    (if (((lhs e_) is_sqrtTerm_in v_) | ((rhs e_) is_sqrtTerm_in v_))" ^
   7.687 +" then (SubProblem (RootEq_,[normalize,root,univariate,equation],          " ^
   7.688 +"       [no_met]) [bool_ e_, real_ v_])                                    " ^
   7.689 +" else (SubProblem (RootEq_,[univariate,equation],                         " ^
   7.690 +"        [no_met]) [bool_ e_, real_ v_]))                                  " ^
   7.691 +" in Check_elementwise L_ {(v_::real). Assumptions})"
   7.692 + ));
   7.693 +
   7.694 +(*-- right 28.08.02 --*)
   7.695 +store_met
   7.696 + (prep_met RootEq.thy "met_rooteq_sq_right" [] e_metID
   7.697 + (["RootEq","solve_right_sq_root_equation"],
   7.698 +   [("#Given" ,["equality e_","solveFor v_"]),
   7.699 +    ("#Where" ,["(rhs e_) is_sqrtTerm_in v_"]),
   7.700 +    ("#Find"  ,["solutions v_i_"])
   7.701 +   ],
   7.702 +   {rew_ord'="termlessI",
   7.703 +    rls'=RootEq_erls,
   7.704 +    srls=e_rls,
   7.705 +    prls=RootEq_prls,
   7.706 +    calc=[],
   7.707 +    crls=RootEq_crls, nrls=norm_Poly},
   7.708 +  "Script Solve_right_sq_root_equation  (e_::bool) (v_::real)  =            " ^
   7.709 +   "(let e_ = " ^
   7.710 +   "    ((Try (Rewrite_Set_Inst [(bdv,v_::real)] r_sqrt_isolate  False)) @@ " ^
   7.711 +   "    (Try (Rewrite_Set                       rooteq_simplify False)) @@  " ^ 
   7.712 +   "    (Try (Repeat (Rewrite_Set expand_rootbinoms            False))) @@  " ^
   7.713 +   "    (Try (Repeat (Rewrite_Set make_rooteq                  False))) @@  " ^
   7.714 +   "    (Try (Rewrite_Set rooteq_simplify                       False))) e_ " ^
   7.715 +   " in if ((rhs e_) is_sqrtTerm_in v_)                                     " ^ 
   7.716 +   " then (SubProblem (RootEq_,[normalize,root,univariate,equation],        " ^
   7.717 +   "       [no_met]) [bool_ e_, real_ v_])                                  " ^
   7.718 +   " else ((SubProblem (RootEq_,[univariate,equation],                      " ^
   7.719 +   "        [no_met]) [bool_ e_, real_ v_])))"
   7.720 + ));
   7.721 +
   7.722 +(*-- left 28.08.02 --*)
   7.723 +store_met
   7.724 + (prep_met RootEq.thy "met_rooteq_sq_left" [] e_metID
   7.725 + (["RootEq","solve_left_sq_root_equation"],
   7.726 +   [("#Given" ,["equality e_","solveFor v_"]),
   7.727 +    ("#Where" ,["(lhs e_) is_sqrtTerm_in v_"]),
   7.728 +    ("#Find"  ,["solutions v_i_"])
   7.729 +   ],
   7.730 +   {rew_ord'="termlessI",
   7.731 +    rls'=RootEq_erls,
   7.732 +    srls=e_rls,
   7.733 +    prls=RootEq_prls,
   7.734 +    calc=[],
   7.735 +    crls=RootEq_crls, nrls=norm_Poly},
   7.736 +    "Script Solve_left_sq_root_equation  (e_::bool) (v_::real)  =          " ^
   7.737 +    "(let e_ =                                                             " ^
   7.738 +    "  ((Try (Rewrite_Set_Inst [(bdv,v_::real)] l_sqrt_isolate  False)) @@ " ^
   7.739 +    "  (Try (Rewrite_Set                       rooteq_simplify False)) @@  " ^
   7.740 +    "  (Try (Repeat (Rewrite_Set expand_rootbinoms            False))) @@  " ^
   7.741 +    "  (Try (Repeat (Rewrite_Set make_rooteq                  False))) @@  " ^
   7.742 +    "  (Try (Rewrite_Set rooteq_simplify                       False))) e_ " ^
   7.743 +    " in if ((lhs e_) is_sqrtTerm_in v_)                                   " ^ 
   7.744 +    " then (SubProblem (RootEq_,[normalize,root,univariate,equation],      " ^
   7.745 +    "       [no_met]) [bool_ e_, real_ v_])                                " ^
   7.746 +    " else ((SubProblem (RootEq_,[univariate,equation],                    " ^
   7.747 +    "        [no_met]) [bool_ e_, real_ v_])))"
   7.748 +   ));
   7.749 +
   7.750 +calclist':= overwritel (!calclist', 
   7.751 +   [("is_rootTerm_in", ("RootEq.is'_rootTerm'_in", 
   7.752 +			eval_is_rootTerm_in"")),
   7.753 +    ("is_sqrtTerm_in", ("RootEq.is'_sqrtTerm'_in", 
   7.754 +			eval_is_sqrtTerm_in"")),
   7.755 +    ("is_normSqrtTerm_in", ("RootEq.is_normSqrtTerm_in", 
   7.756 +				 eval_is_normSqrtTerm_in""))
   7.757 +    ]);(*("", ("", "")),*)
   7.758 +*}
   7.759 +
   7.760  end
     8.1 --- a/src/Tools/isac/Knowledge/Simplify.thy	Thu Aug 26 10:03:53 2010 +0200
     8.2 +++ b/src/Tools/isac/Knowledge/Simplify.thy	Thu Aug 26 18:15:30 2010 +0200
     8.3 @@ -9,13 +9,13 @@
     8.4  use_thy"~/proto2/isac/src/sml/Knowledge/Isac";
     8.5  *)
     8.6  
     8.7 -Simplify = Atools +
     8.8 +theory Simplify imports Atools begin
     8.9  
    8.10  consts
    8.11  
    8.12    (*descriptions in the related problem*)
    8.13 -  term        :: real => una
    8.14 -  normalform  :: real => una
    8.15 +  term        :: "real => una"
    8.16 +  normalform  :: "real => una"
    8.17  
    8.18    (*the CAS-command*)
    8.19    Simplify    :: "real => real"  (*"Simplify (1+2a+3+4a)*)
    8.20 @@ -25,5 +25,70 @@
    8.21    SimplifyScript      :: "[real,  real] => real"
    8.22                    ("((Script SimplifyScript (_ =))// (_))" 9)
    8.23  
    8.24 +ML {*
    8.25 +
    8.26 +(** problems **)
    8.27 +store_pbt
    8.28 + (prep_pbt Simplify.thy "pbl_simp" [] e_pblID
    8.29 + (["simplification"],
    8.30 +  [("#Given" ,["term t_"]),
    8.31 +   ("#Find"  ,["normalform n_"])
    8.32 +  ],
    8.33 +  append_rls "e_rls" e_rls [(*for preds in where_*)], 
    8.34 +  SOME "Simplify t_", 
    8.35 +  []));
    8.36 +
    8.37 +store_pbt
    8.38 + (prep_pbt Simplify.thy "pbl_vereinfache" [] e_pblID
    8.39 + (["vereinfachen"],
    8.40 +  [("#Given" ,["term t_"]),
    8.41 +   ("#Find"  ,["normalform n_"])
    8.42 +  ],
    8.43 +  append_rls "e_rls" e_rls [(*for preds in where_*)], 
    8.44 +  SOME "Vereinfache t_", 
    8.45 +  []));
    8.46 +
    8.47 +(** methods **)
    8.48 +
    8.49 +store_met
    8.50 +    (prep_met Simplify.thy "met_simp" [] e_metID
    8.51 +	      (["simplification"],
    8.52 +	       [("#Given" ,["term t_"]),
    8.53 +		("#Find"  ,["normalform n_"])
    8.54 +		],
    8.55 +	       {rew_ord'="tless_true",
    8.56 +		rls'= e_rls, 
    8.57 +		calc = [], 
    8.58 +		srls = e_rls, 
    8.59 +		prls=e_rls,
    8.60 +		crls = e_rls, nrls = e_rls},
    8.61 +	       "empty_script"
    8.62 +	       ));
    8.63 +
    8.64 +(** CAS-command **)
    8.65 +
    8.66 +(*.function for handling the cas-input "Simplify (2*a + 3*a)":
    8.67 +   make a model which is already in ptree-internal format.*)
    8.68 +(* val (h,argl) = strip_comb (str2term "Simplify (2*a + 3*a)");
    8.69 +   val (h,argl) = strip_comb ((term_of o the o (parse thy)) 
    8.70 +				  "Simplify (2*a + 3*a)");
    8.71 +   *)
    8.72 +fun argl2dtss t =
    8.73 +    [((term_of o the o (parse thy)) "term", t),
    8.74 +     ((term_of o the o (parse thy)) "normalform", 
    8.75 +      [(term_of o the o (parse thy)) "N"])
    8.76 +     ]
    8.77 +  | argl2dtss _ = raise error "Simplify.ML: wrong argument for argl2dtss";
    8.78 +
    8.79 +castab := 
    8.80 +overwritel (!castab, 
    8.81 +	    [((term_of o the o (parse thy)) "Simplify",  
    8.82 +	      (("Isac.thy", ["simplification"], ["no_met"]), 
    8.83 +	       argl2dtss)),
    8.84 +	     ((term_of o the o (parse thy)) "Vereinfache",  
    8.85 +	      (("Isac.thy", ["vereinfachen"], ["no_met"]), 
    8.86 +	       argl2dtss))
    8.87 +	     ]);
    8.88 +*}
    8.89  
    8.90  end
    8.91 \ No newline at end of file