src/Tools/isac/Knowledge/RootEq.thy
author Walther Neuper <neuper@ist.tugraz.at>
Sun, 22 Sep 2013 18:09:05 +0200
changeset 52125 6f1d3415dc68
parent 48789 498ed5bb1004
child 52145 d13173b915e0
permissions -rw-r--r--
add functions accessing Theory_Data in parallel to those accessing "ruleset' = Unsynchronized.ref"

updates have been done incrementally following Build_Isac.thy:
# ./bin/isabelle jedit -l HOL src/Tools/isac/ProgLang/ProgLang.thy &
# ./bin/isabelle jedit -l HOL src/Tools/isac/Interpret/Interpret.thy &
# ./bin/isabelle jedit -l HOL src/Tools/isac/xmlsrc/xmlsrc.thy &
# ./bin/isabelle jedit -l HOL src/Tools/isac/Frontend/Frontend.thy &

Note, that the original access function "fun assoc_rls" is still outcommented;
so the old and new functionality is established in parallel.
     1 (*.(c) by Richard Lang, 2003 .*)
     2 (* collecting all knowledge for Root Equations
     3    created by: rlang 
     4          date: 02.08
     5    changed by: rlang
     6    last change by: rlang
     7              date: 02.11.14
     8 *)
     9 
    10 theory RootEq imports Root begin
    11 
    12 text {* univariate equations containing real square roots:
    13   This type of equations has been used to bootstrap Lucas-Interpretation.
    14   In 2003 this type has been extended and integrated into ISAC's equation solver
    15   by Richard Lang in 2003.
    16   The assumptions (all containing "<") didn't pass the xml-parsers at the 
    17   interface between math-engine and front-end.
    18   The migration Isabelle2002 --> 2011 dropped this type of equation, see
    19   test/../rooteq.sml, rootrateq.sml.
    20 *}
    21 
    22 consts
    23 
    24   is'_rootTerm'_in     :: "[real, real] => bool" ("_ is'_rootTerm'_in _")
    25   is'_sqrtTerm'_in     :: "[real, real] => bool" ("_ is'_sqrtTerm'_in _") 
    26   is'_normSqrtTerm'_in :: "[real, real] => bool" ("_ is'_normSqrtTerm'_in _") 
    27 
    28   (*----------------------scripts-----------------------*)
    29   Norm'_sq'_root'_equation
    30              :: "[bool,real, 
    31 		   bool list] => bool list"
    32                ("((Script Norm'_sq'_root'_equation (_ _ =))// 
    33                   (_))" 9)
    34   Solve'_sq'_root'_equation
    35              :: "[bool,real, 
    36 		   bool list] => bool list"
    37                ("((Script Solve'_sq'_root'_equation (_ _ =))// 
    38                   (_))" 9)
    39   Solve'_left'_sq'_root'_equation
    40              :: "[bool,real, 
    41 		   bool list] => bool list"
    42                ("((Script Solve'_left'_sq'_root'_equation (_ _ =))// 
    43                   (_))" 9)
    44   Solve'_right'_sq'_root'_equation
    45              :: "[bool,real, 
    46 		   bool list] => bool list"
    47                ("((Script Solve'_right'_sq'_root'_equation (_ _ =))// 
    48                   (_))" 9)
    49  
    50 axioms(*axiomatization where*)
    51 
    52 (* normalize *)
    53   makex1_x:            "a^^^1  = a"   (*and*)
    54   real_assoc_1:        "a+(b+c) = a+b+c" (*and*)
    55   real_assoc_2:        "a*(b*c) = a*b*c" (*and*)
    56 
    57   (* simplification of root*)
    58   sqrt_square_1:       "[|0 <= a|] ==>  (sqrt a)^^^2 = a" (*and*)
    59   sqrt_square_2:       "sqrt (a ^^^ 2) = a" (*and*)
    60   sqrt_times_root_1:   "sqrt a * sqrt b = sqrt(a*b)" (*and*)
    61   sqrt_times_root_2:   "a * sqrt b * sqrt c = a * sqrt(b*c)" (*and*)
    62 
    63   (* isolate one root on the LEFT or RIGHT hand side of the equation *)
    64   sqrt_isolate_l_add1: "[|bdv occurs_in c|] ==> 
    65    (a + b*sqrt(c) = d) = (b * sqrt(c) = d+ (-1) * a)" (*and*)
    66   sqrt_isolate_l_add2: "[|bdv occurs_in c|] ==>
    67    (a + sqrt(c) = d) = ((sqrt(c) = d+ (-1) * a))" (*and*)
    68   sqrt_isolate_l_add3: "[|bdv occurs_in c|] ==>
    69    (a + b*(e/sqrt(c)) = d) = (b * (e/sqrt(c)) = d + (-1) * a)" (*and*)
    70   sqrt_isolate_l_add4: "[|bdv occurs_in c|] ==>
    71    (a + b/(f*sqrt(c)) = d) = (b / (f*sqrt(c)) = d + (-1) * a)" (*and*)
    72   sqrt_isolate_l_add5: "[|bdv occurs_in c|] ==>
    73    (a + b*(e/(f*sqrt(c))) = d) = (b * (e/(f*sqrt(c))) = d+ (-1) * a)" (*and*)
    74   sqrt_isolate_l_add6: "[|bdv occurs_in c|] ==>
    75    (a + b/sqrt(c) = d) = (b / sqrt(c) = d+ (-1) * a)" (*and*)
    76   sqrt_isolate_r_add1: "[|bdv occurs_in f|] ==>
    77    (a = d + e*sqrt(f)) = (a + (-1) * d = e*sqrt(f))" (*and*)
    78   sqrt_isolate_r_add2: "[|bdv occurs_in f|] ==>
    79    (a = d + sqrt(f)) = (a + (-1) * d = sqrt(f))" (*and*)
    80  (* small hack: thm 3,5,6 are not needed if rootnormalize is well done*)
    81   sqrt_isolate_r_add3: "[|bdv occurs_in f|] ==>
    82    (a = d + e*(g/sqrt(f))) = (a + (-1) * d = e*(g/sqrt(f)))" (*and*)
    83   sqrt_isolate_r_add4: "[|bdv occurs_in f|] ==>
    84    (a = d + g/sqrt(f)) = (a + (-1) * d = g/sqrt(f))" (*and*)
    85   sqrt_isolate_r_add5: "[|bdv occurs_in f|] ==>
    86    (a = d + e*(g/(h*sqrt(f)))) = (a + (-1) * d = e*(g/(h*sqrt(f))))" (*and*)
    87   sqrt_isolate_r_add6: "[|bdv occurs_in f|] ==>
    88    (a = d + g/(h*sqrt(f))) = (a + (-1) * d = g/(h*sqrt(f)))" (*and*)
    89  
    90   (* eliminate isolates sqrt *)
    91   sqrt_square_equation_both_1: "[|bdv occurs_in b; bdv occurs_in d|] ==> 
    92    ( (sqrt a + sqrt b         = sqrt c + sqrt d) = 
    93      (a+2*sqrt(a)*sqrt(b)+b  = c+2*sqrt(c)*sqrt(d)+d))" (*and*)
    94   sqrt_square_equation_both_2: "[|bdv occurs_in b; bdv occurs_in d|] ==> 
    95    ( (sqrt a - sqrt b           = sqrt c + sqrt d) = 
    96      (a - 2*sqrt(a)*sqrt(b)+b  = c+2*sqrt(c)*sqrt(d)+d))" (*and*)
    97   sqrt_square_equation_both_3: "[|bdv occurs_in b; bdv occurs_in d|] ==> 
    98    ( (sqrt a + sqrt b           = sqrt c - sqrt d) = 
    99      (a + 2*sqrt(a)*sqrt(b)+b  = c - 2*sqrt(c)*sqrt(d)+d))" (*and*)
   100   sqrt_square_equation_both_4: "[|bdv occurs_in b; bdv occurs_in d|] ==> 
   101    ( (sqrt a - sqrt b           = sqrt c - sqrt d) = 
   102      (a - 2*sqrt(a)*sqrt(b)+b  = c - 2*sqrt(c)*sqrt(d)+d))" (*and*)
   103   sqrt_square_equation_left_1: "[|bdv occurs_in a; 0 <= a; 0 <= b|] ==>
   104    ( (sqrt (a) = b) = (a = (b^^^2)))" (*and*)
   105   sqrt_square_equation_left_2: "[|bdv occurs_in a; 0 <= a; 0 <= b*c|] ==> 
   106    ( (c*sqrt(a) = b) = (c^^^2*a = b^^^2))" (*and*)
   107   sqrt_square_equation_left_3: "[|bdv occurs_in a; 0 <= a; 0 <= b*c|] ==> 
   108    ( c/sqrt(a) = b) = (c^^^2 / a = b^^^2)" (*and*)
   109   (* small hack: thm 4-6 are not needed if rootnormalize is well done*)
   110   sqrt_square_equation_left_4: "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d|] ==> 
   111    ( (c*(d/sqrt (a)) = b) = (c^^^2*(d^^^2/a) = b^^^2))" (*and*)
   112   sqrt_square_equation_left_5: "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d|] ==> 
   113    ( c/(d*sqrt(a)) = b) = (c^^^2 / (d^^^2*a) = b^^^2)" (*and*)
   114   sqrt_square_equation_left_6: "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d*e|] ==> 
   115    ( (c*(d/(e*sqrt (a))) = b) = (c^^^2*(d^^^2/(e^^^2*a)) = b^^^2))" (*and*)
   116   sqrt_square_equation_right_1:  "[|bdv occurs_in b; 0 <= a; 0 <= b|] ==> 
   117    ( (a = sqrt (b)) = (a^^^2 = b))" (*and*)
   118   sqrt_square_equation_right_2: "[|bdv occurs_in b; 0 <= a*c; 0 <= b|] ==> 
   119    ( (a = c*sqrt (b)) = ((a^^^2) = c^^^2*b))" (*and*)
   120   sqrt_square_equation_right_3: "[|bdv occurs_in b; 0 <= a*c; 0 <= b|] ==> 
   121    ( (a = c/sqrt (b)) = (a^^^2 = c^^^2/b))" (*and*)
   122  (* small hack: thm 4-6 are not needed if rootnormalize is well done*)
   123   sqrt_square_equation_right_4: "[|bdv occurs_in b; 0 <= a*c*d; 0 <= b|] ==> 
   124    ( (a = c*(d/sqrt (b))) = ((a^^^2) = c^^^2*(d^^^2/b)))" (*and*)
   125   sqrt_square_equation_right_5: "[|bdv occurs_in b; 0 <= a*c*d; 0 <= b|] ==> 
   126    ( (a = c/(d*sqrt (b))) = (a^^^2 = c^^^2/(d^^^2*b)))" (*and*)
   127   sqrt_square_equation_right_6: "[|bdv occurs_in b; 0 <= a*c*d*e; 0 <= b|] ==> 
   128    ( (a = c*(d/(e*sqrt (b)))) = ((a^^^2) = c^^^2*(d^^^2/(e^^^2*b))))"
   129 
   130 ML {*
   131 val thy = @{theory};
   132 
   133 (*-------------------------functions---------------------*)
   134 (* true if bdv is under sqrt of a Equation*)
   135 fun is_rootTerm_in t v = 
   136     let 
   137 	fun coeff_in c v = member op = (vars c) v;
   138    	fun findroot (_ $ _ $ _ $ _) v = error("is_rootTerm_in:")
   139 	  (* at the moment there is no term like this, but ....*)
   140 	  | findroot (t as (Const ("Root.nroot",_) $ _ $ t3)) v = coeff_in t3 v
   141 	  | findroot (_ $ t2 $ t3) v = (findroot t2 v) orelse (findroot t3 v)
   142 	  | findroot (t as (Const ("NthRoot.sqrt",_) $ t2)) v = coeff_in t2 v
   143 	  | findroot (_ $ t2) v = (findroot t2 v)
   144 	  | findroot _ _ = false;
   145      in
   146 	findroot t v
   147     end;
   148 
   149  fun is_sqrtTerm_in t v = 
   150     let 
   151 	fun coeff_in c v = member op = (vars c) v;
   152    	fun findsqrt (_ $ _ $ _ $ _) v = error("is_sqrteqation_in:")
   153 	  (* at the moment there is no term like this, but ....*)
   154 	  | findsqrt (_ $ t1 $ t2) v = (findsqrt t1 v) orelse (findsqrt t2 v)
   155 	  | findsqrt (t as (Const ("NthRoot.sqrt",_) $ a)) v = coeff_in a v
   156 	  | findsqrt (_ $ t1) v = (findsqrt t1 v)
   157 	  | findsqrt _ _ = false;
   158      in
   159 	findsqrt t v
   160     end;
   161 
   162 (* RL: 030518: Is in the rightest subterm of a term a sqrt with bdv,
   163 and the subterm ist connected with + or * --> is normalized*)
   164  fun is_normSqrtTerm_in t v =
   165      let
   166 	fun coeff_in c v = member op = (vars c) v;
   167         fun isnorm (_ $ _ $ _ $ _) v = error("is_normSqrtTerm_in:")
   168 	  (* at the moment there is no term like this, but ....*)
   169           | isnorm (Const ("Groups.plus_class.plus",_) $ _ $ t2) v = is_sqrtTerm_in t2 v
   170           | isnorm (Const ("Groups.times_class.times",_) $ _ $ t2) v = is_sqrtTerm_in t2 v
   171           | isnorm (Const ("Groups.minus_class.minus",_) $ _ $ _) v = false
   172           | isnorm (Const ("Fields.inverse_class.divide",_) $ t1 $ t2) v = (is_sqrtTerm_in t1 v) orelse 
   173                               (is_sqrtTerm_in t2 v)
   174           | isnorm (Const ("NthRoot.sqrt",_) $ t1) v = coeff_in t1 v
   175  	  | isnorm (_ $ t1) v = is_sqrtTerm_in t1 v
   176           | isnorm _ _ = false;
   177      in
   178          isnorm t v
   179      end;
   180 
   181 fun eval_is_rootTerm_in _ _ 
   182        (p as (Const ("RootEq.is'_rootTerm'_in",_) $ t $ v)) _  =
   183     if is_rootTerm_in t v then 
   184 	SOME ((term2str p) ^ " = True",
   185 	      Trueprop $ (mk_equality (p, @{term True})))
   186     else SOME ((term2str p) ^ " = True",
   187 	       Trueprop $ (mk_equality (p, @{term False})))
   188   | eval_is_rootTerm_in _ _ _ _ = ((*tracing"### nichts matcht";*) NONE);
   189 
   190 fun eval_is_sqrtTerm_in _ _ 
   191        (p as (Const ("RootEq.is'_sqrtTerm'_in",_) $ t $ v)) _  =
   192     if is_sqrtTerm_in t v then 
   193 	SOME ((term2str p) ^ " = True",
   194 	      Trueprop $ (mk_equality (p, @{term True})))
   195     else SOME ((term2str p) ^ " = True",
   196 	       Trueprop $ (mk_equality (p, @{term False})))
   197   | eval_is_sqrtTerm_in _ _ _ _ = ((*tracing"### nichts matcht";*) NONE);
   198 
   199 fun eval_is_normSqrtTerm_in _ _ 
   200        (p as (Const ("RootEq.is'_normSqrtTerm'_in",_) $ t $ v)) _  =
   201     if is_normSqrtTerm_in t v then 
   202 	SOME ((term2str p) ^ " = True",
   203 	      Trueprop $ (mk_equality (p, @{term True})))
   204     else SOME ((term2str p) ^ " = True",
   205 	       Trueprop $ (mk_equality (p, @{term False})))
   206   | eval_is_normSqrtTerm_in _ _ _ _ = ((*tracing"### nichts matcht";*) NONE);
   207 
   208 (*-------------------------rulse-------------------------*)
   209 val RootEq_prls =(*15.10.02:just the following order due to subterm evaluation*)
   210   append_rls "RootEq_prls" e_rls 
   211 	     [Calc ("Atools.ident",eval_ident "#ident_"),
   212 	      Calc ("Tools.matches",eval_matches ""),
   213 	      Calc ("Tools.lhs"    ,eval_lhs ""),
   214 	      Calc ("Tools.rhs"    ,eval_rhs ""),
   215 	      Calc ("RootEq.is'_sqrtTerm'_in",eval_is_sqrtTerm_in ""),
   216 	      Calc ("RootEq.is'_rootTerm'_in",eval_is_rootTerm_in ""),
   217 	      Calc ("RootEq.is'_normSqrtTerm'_in",eval_is_normSqrtTerm_in ""),
   218 	      Calc ("HOL.eq",eval_equal "#equal_"),
   219 	      Thm ("not_true",num_str @{thm not_true}),
   220 	      Thm ("not_false",num_str @{thm not_false}),
   221 	      Thm ("and_true",num_str @{thm and_true}),
   222 	      Thm ("and_false",num_str @{thm and_false}),
   223 	      Thm ("or_true",num_str @{thm or_true}),
   224 	      Thm ("or_false",num_str @{thm or_false})
   225 	      ];
   226 
   227 val RootEq_erls =
   228      append_rls "RootEq_erls" Root_erls
   229           [Thm ("divide_divide_eq_left",num_str @{thm divide_divide_eq_left})
   230            ];
   231 
   232 val RootEq_crls = 
   233      append_rls "RootEq_crls" Root_crls
   234           [Thm ("divide_divide_eq_left",num_str @{thm divide_divide_eq_left})
   235            ];
   236 
   237 val rooteq_srls = 
   238      append_rls "rooteq_srls" e_rls
   239 		[Calc ("RootEq.is'_sqrtTerm'_in",eval_is_sqrtTerm_in ""),
   240                  Calc ("RootEq.is'_normSqrtTerm'_in",eval_is_normSqrtTerm_in""),
   241                  Calc ("RootEq.is'_rootTerm'_in",eval_is_rootTerm_in "")
   242 		 ];
   243 
   244 ruleset' := overwritelthy @{theory} (!ruleset',
   245 			[("RootEq_erls",RootEq_erls),
   246                                              (*FIXXXME:del with rls.rls'*)
   247 			 ("rooteq_srls",rooteq_srls)
   248                          ]);
   249 *}
   250 setup {* KEStore_Elems.add_rlss 
   251   [("RootEq_erls", (Context.theory_name @{theory}, RootEq_erls)), 
   252   ("rooteq_srls", (Context.theory_name @{theory}, rooteq_srls))] *}
   253 ML {*
   254 
   255 (*isolate the bound variable in an sqrt equation; 'bdv' is a meta-constant*)
   256  val sqrt_isolate = prep_rls(
   257   Rls {id = "sqrt_isolate", preconds = [], rew_ord = ("termlessI",termlessI), 
   258        erls = RootEq_erls, srls = Erls, calc = [], errpatts = [],
   259        rules = [
   260        Thm("sqrt_square_1",num_str @{thm sqrt_square_1}),
   261                      (* (sqrt a)^^^2 -> a *)
   262        Thm("sqrt_square_2",num_str @{thm sqrt_square_2}),
   263                      (* sqrt (a^^^2) -> a *)
   264        Thm("sqrt_times_root_1",num_str @{thm sqrt_times_root_1}),
   265             (* sqrt a sqrt b -> sqrt(ab) *)
   266        Thm("sqrt_times_root_2",num_str @{thm sqrt_times_root_2}),
   267             (* a sqrt b sqrt c -> a sqrt(bc) *)
   268        Thm("sqrt_square_equation_both_1",
   269            num_str @{thm sqrt_square_equation_both_1}),
   270        (* (sqrt a + sqrt b  = sqrt c + sqrt d) -> 
   271             (a+2*sqrt(a)*sqrt(b)+b) = c+2*sqrt(c)*sqrt(d)+d) *)
   272        Thm("sqrt_square_equation_both_2",
   273             num_str @{thm sqrt_square_equation_both_2}),
   274        (* (sqrt a - sqrt b  = sqrt c + sqrt d) -> 
   275             (a-2*sqrt(a)*sqrt(b)+b) = c+2*sqrt(c)*sqrt(d)+d) *)
   276        Thm("sqrt_square_equation_both_3",
   277             num_str @{thm sqrt_square_equation_both_3}),
   278        (* (sqrt a + sqrt b  = sqrt c - sqrt d) -> 
   279             (a+2*sqrt(a)*sqrt(b)+b) = c-2*sqrt(c)*sqrt(d)+d) *)
   280        Thm("sqrt_square_equation_both_4",
   281             num_str @{thm sqrt_square_equation_both_4}),
   282        (* (sqrt a - sqrt b  = sqrt c - sqrt d) -> 
   283             (a-2*sqrt(a)*sqrt(b)+b) = c-2*sqrt(c)*sqrt(d)+d) *)
   284        Thm("sqrt_isolate_l_add1",
   285             num_str @{thm sqrt_isolate_l_add1}), 
   286        (* a+b*sqrt(x)=d -> b*sqrt(x) = d-a *)
   287        Thm("sqrt_isolate_l_add2",
   288             num_str @{thm sqrt_isolate_l_add2}), 
   289        (* a+  sqrt(x)=d ->   sqrt(x) = d-a *)
   290        Thm("sqrt_isolate_l_add3",
   291             num_str @{thm sqrt_isolate_l_add3}), 
   292        (* a+b*c/sqrt(x)=d->b*c/sqrt(x)=d-a *)
   293        Thm("sqrt_isolate_l_add4",
   294             num_str @{thm sqrt_isolate_l_add4}), 
   295        (* a+c/sqrt(x)=d -> c/sqrt(x) = d-a *)
   296        Thm("sqrt_isolate_l_add5",
   297             num_str @{thm sqrt_isolate_l_add5}), 
   298        (* a+b*c/f*sqrt(x)=d->b*c/f*sqrt(x)=d-a *)
   299        Thm("sqrt_isolate_l_add6",
   300             num_str @{thm sqrt_isolate_l_add6}), 
   301        (* a+c/f*sqrt(x)=d -> c/f*sqrt(x) = d-a *)
   302        (*Thm("sqrt_isolate_l_div",num_str @{thm sqrt_isolate_l_div}),*)
   303          (* b*sqrt(x) = d sqrt(x) d/b *)
   304        Thm("sqrt_isolate_r_add1",
   305             num_str @{thm sqrt_isolate_r_add1}),
   306        (* a= d+e*sqrt(x) -> a-d=e*sqrt(x) *)
   307        Thm("sqrt_isolate_r_add2",
   308             num_str @{thm sqrt_isolate_r_add2}),
   309        (* a= d+  sqrt(x) -> a-d=  sqrt(x) *)
   310        Thm("sqrt_isolate_r_add3",
   311             num_str @{thm sqrt_isolate_r_add3}),
   312        (* a=d+e*g/sqrt(x)->a-d=e*g/sqrt(x)*)
   313        Thm("sqrt_isolate_r_add4",
   314             num_str @{thm sqrt_isolate_r_add4}),
   315        (* a= d+g/sqrt(x) -> a-d=g/sqrt(x) *)
   316        Thm("sqrt_isolate_r_add5",
   317             num_str @{thm sqrt_isolate_r_add5}),
   318        (* a=d+e*g/h*sqrt(x)->a-d=e*g/h*sqrt(x)*)
   319        Thm("sqrt_isolate_r_add6",
   320             num_str @{thm sqrt_isolate_r_add6}),
   321        (* a= d+g/h*sqrt(x) -> a-d=g/h*sqrt(x) *)
   322        (*Thm("sqrt_isolate_r_div",num_str @{thm sqrt_isolate_r_div}),*)
   323          (* a=e*sqrt(x) -> a/e = sqrt(x) *)
   324        Thm("sqrt_square_equation_left_1",
   325             num_str @{thm sqrt_square_equation_left_1}),   
   326        (* sqrt(x)=b -> x=b^2 *)
   327        Thm("sqrt_square_equation_left_2",
   328             num_str @{thm sqrt_square_equation_left_2}),   
   329        (* c*sqrt(x)=b -> c^2*x=b^2 *)
   330        Thm("sqrt_square_equation_left_3",num_str @{thm sqrt_square_equation_left_3}),  
   331 	      (* c/sqrt(x)=b -> c^2/x=b^2 *)
   332        Thm("sqrt_square_equation_left_4",num_str @{thm sqrt_square_equation_left_4}),
   333 	      (* c*d/sqrt(x)=b -> c^2*d^2/x=b^2 *)
   334        Thm("sqrt_square_equation_left_5",num_str @{thm sqrt_square_equation_left_5}),
   335 	      (* c/d*sqrt(x)=b -> c^2/d^2x=b^2 *)
   336        Thm("sqrt_square_equation_left_6",num_str @{thm sqrt_square_equation_left_6}),
   337 	      (* c*d/g*sqrt(x)=b -> c^2*d^2/g^2x=b^2 *)
   338        Thm("sqrt_square_equation_right_1",num_str @{thm sqrt_square_equation_right_1}),
   339 	      (* a=sqrt(x) ->a^2=x *)
   340        Thm("sqrt_square_equation_right_2",num_str @{thm sqrt_square_equation_right_2}),
   341 	      (* a=c*sqrt(x) ->a^2=c^2*x *)
   342        Thm("sqrt_square_equation_right_3",num_str @{thm sqrt_square_equation_right_3}),
   343 	      (* a=c/sqrt(x) ->a^2=c^2/x *)
   344        Thm("sqrt_square_equation_right_4",num_str @{thm sqrt_square_equation_right_4}),
   345 	      (* a=c*d/sqrt(x) ->a^2=c^2*d^2/x *)
   346        Thm("sqrt_square_equation_right_5",num_str @{thm sqrt_square_equation_right_5}),
   347 	      (* a=c/e*sqrt(x) ->a^2=c^2/e^2x *)
   348        Thm("sqrt_square_equation_right_6",num_str @{thm sqrt_square_equation_right_6})
   349 	      (* a=c*d/g*sqrt(x) ->a^2=c^2*d^2/g^2*x *)
   350        ],scr = Prog ((term_of o the o (parse thy)) "empty_script")
   351       }:rls);
   352 
   353 ruleset' := overwritelthy @{theory} (!ruleset',
   354 			[("sqrt_isolate",sqrt_isolate)
   355 			 ]);
   356 *}
   357 setup {* KEStore_Elems.add_rlss
   358   [("sqrt_isolate", (Context.theory_name @{theory}, sqrt_isolate))] *}
   359 ML {*
   360 
   361 (*isolate the bound variable in an sqrt left equation; 'bdv' is a meta-constant*)
   362  val l_sqrt_isolate = prep_rls(
   363      Rls {id = "l_sqrt_isolate", preconds = [], 
   364 	  rew_ord = ("termlessI",termlessI), 
   365           erls = RootEq_erls, srls = Erls, calc = [], errpatts = [],
   366      rules = [
   367      Thm("sqrt_square_1",num_str @{thm sqrt_square_1}),
   368                             (* (sqrt a)^^^2 -> a *)
   369      Thm("sqrt_square_2",num_str @{thm sqrt_square_2}),
   370                             (* sqrt (a^^^2) -> a *)
   371      Thm("sqrt_times_root_1",num_str @{thm sqrt_times_root_1}),
   372             (* sqrt a sqrt b -> sqrt(ab) *)
   373      Thm("sqrt_times_root_2",num_str @{thm sqrt_times_root_2}),
   374         (* a sqrt b sqrt c -> a sqrt(bc) *)
   375      Thm("sqrt_isolate_l_add1",num_str @{thm sqrt_isolate_l_add1}),
   376         (* a+b*sqrt(x)=d -> b*sqrt(x) = d-a *)
   377      Thm("sqrt_isolate_l_add2",num_str @{thm sqrt_isolate_l_add2}),
   378         (* a+  sqrt(x)=d ->   sqrt(x) = d-a *)
   379      Thm("sqrt_isolate_l_add3",num_str @{thm sqrt_isolate_l_add3}),
   380         (* a+b*c/sqrt(x)=d->b*c/sqrt(x)=d-a *)
   381      Thm("sqrt_isolate_l_add4",num_str @{thm sqrt_isolate_l_add4}),
   382         (* a+c/sqrt(x)=d -> c/sqrt(x) = d-a *)
   383      Thm("sqrt_isolate_l_add5",num_str @{thm sqrt_isolate_l_add5}),
   384         (* a+b*c/f*sqrt(x)=d->b*c/f*sqrt(x)=d-a *)
   385      Thm("sqrt_isolate_l_add6",num_str @{thm sqrt_isolate_l_add6}),
   386         (* a+c/f*sqrt(x)=d -> c/f*sqrt(x) = d-a *)
   387    (*Thm("sqrt_isolate_l_div",num_str @{thm sqrt_isolate_l_div}),*)
   388         (* b*sqrt(x) = d sqrt(x) d/b *)
   389      Thm("sqrt_square_equation_left_1",num_str @{thm sqrt_square_equation_left_1}),
   390 	      (* sqrt(x)=b -> x=b^2 *)
   391      Thm("sqrt_square_equation_left_2",num_str @{thm sqrt_square_equation_left_2}),
   392 	      (* a*sqrt(x)=b -> a^2*x=b^2*)
   393      Thm("sqrt_square_equation_left_3",num_str @{thm sqrt_square_equation_left_3}),   
   394 	      (* c/sqrt(x)=b -> c^2/x=b^2 *)
   395      Thm("sqrt_square_equation_left_4",num_str @{thm sqrt_square_equation_left_4}),   
   396 	      (* c*d/sqrt(x)=b -> c^2*d^2/x=b^2 *)
   397      Thm("sqrt_square_equation_left_5",num_str @{thm sqrt_square_equation_left_5}),   
   398 	      (* c/d*sqrt(x)=b -> c^2/d^2x=b^2 *)
   399      Thm("sqrt_square_equation_left_6",num_str @{thm sqrt_square_equation_left_6})  
   400 	      (* c*d/g*sqrt(x)=b -> c^2*d^2/g^2x=b^2 *)
   401     ],
   402     scr = Prog ((term_of o the o (parse thy)) "empty_script")
   403    }:rls);
   404 
   405 ruleset' := overwritelthy @{theory} (!ruleset',
   406 			[("l_sqrt_isolate",l_sqrt_isolate)
   407 			 ]);
   408 *}
   409 setup {* KEStore_Elems.add_rlss
   410   [("l_sqrt_isolate", (Context.theory_name @{theory}, l_sqrt_isolate))] *}
   411 ML {*
   412 
   413 (* -- right 28.8.02--*)
   414 (*isolate the bound variable in an sqrt right equation; 'bdv' is a meta-constant*)
   415  val r_sqrt_isolate = prep_rls(
   416      Rls {id = "r_sqrt_isolate", preconds = [], 
   417 	  rew_ord = ("termlessI",termlessI), 
   418           erls = RootEq_erls, srls = Erls, calc = [], errpatts = [],
   419      rules = [
   420      Thm("sqrt_square_1",num_str @{thm sqrt_square_1}),
   421                            (* (sqrt a)^^^2 -> a *)
   422      Thm("sqrt_square_2",num_str @{thm sqrt_square_2}), 
   423                            (* sqrt (a^^^2) -> a *)
   424      Thm("sqrt_times_root_1",num_str @{thm sqrt_times_root_1}),
   425            (* sqrt a sqrt b -> sqrt(ab) *)
   426      Thm("sqrt_times_root_2",num_str @{thm sqrt_times_root_2}),
   427        (* a sqrt b sqrt c -> a sqrt(bc) *)
   428      Thm("sqrt_isolate_r_add1",num_str @{thm sqrt_isolate_r_add1}),
   429        (* a= d+e*sqrt(x) -> a-d=e*sqrt(x) *)
   430      Thm("sqrt_isolate_r_add2",num_str @{thm sqrt_isolate_r_add2}),
   431        (* a= d+  sqrt(x) -> a-d=  sqrt(x) *)
   432      Thm("sqrt_isolate_r_add3",num_str @{thm sqrt_isolate_r_add3}),
   433        (* a=d+e*g/sqrt(x)->a-d=e*g/sqrt(x)*)
   434      Thm("sqrt_isolate_r_add4",num_str @{thm sqrt_isolate_r_add4}),
   435        (* a= d+g/sqrt(x) -> a-d=g/sqrt(x) *)
   436      Thm("sqrt_isolate_r_add5",num_str @{thm sqrt_isolate_r_add5}),
   437        (* a=d+e*g/h*sqrt(x)->a-d=e*g/h*sqrt(x)*)
   438      Thm("sqrt_isolate_r_add6",num_str @{thm sqrt_isolate_r_add6}),
   439        (* a= d+g/h*sqrt(x) -> a-d=g/h*sqrt(x) *)
   440    (*Thm("sqrt_isolate_r_div",num_str @{thm sqrt_isolate_r_div}),*)
   441        (* a=e*sqrt(x) -> a/e = sqrt(x) *)
   442      Thm("sqrt_square_equation_right_1",num_str @{thm sqrt_square_equation_right_1}),
   443 	      (* a=sqrt(x) ->a^2=x *)
   444      Thm("sqrt_square_equation_right_2",num_str @{thm sqrt_square_equation_right_2}),
   445 	      (* a=c*sqrt(x) ->a^2=c^2*x *)
   446      Thm("sqrt_square_equation_right_3",num_str @{thm sqrt_square_equation_right_3}),
   447 	      (* a=c/sqrt(x) ->a^2=c^2/x *)
   448      Thm("sqrt_square_equation_right_4",num_str @{thm sqrt_square_equation_right_4}), 
   449 	      (* a=c*d/sqrt(x) ->a^2=c^2*d^2/x *)
   450      Thm("sqrt_square_equation_right_5",num_str @{thm sqrt_square_equation_right_5}),
   451 	      (* a=c/e*sqrt(x) ->a^2=c^2/e^2x *)
   452      Thm("sqrt_square_equation_right_6",num_str @{thm sqrt_square_equation_right_6})
   453 	      (* a=c*d/g*sqrt(x) ->a^2=c^2*d^2/g^2*x *)
   454     ],
   455     scr = Prog ((term_of o the o (parse thy)) "empty_script")
   456    }:rls);
   457 
   458 ruleset' := overwritelthy @{theory} (!ruleset',
   459 			[("r_sqrt_isolate",r_sqrt_isolate)
   460 			 ]);
   461 *}
   462 setup {* KEStore_Elems.add_rlss
   463   [("r_sqrt_isolate", (Context.theory_name @{theory}, r_sqrt_isolate))] *}
   464 ML {*
   465 
   466 val rooteq_simplify = prep_rls(
   467   Rls {id = "rooteq_simplify", 
   468        preconds = [], rew_ord = ("termlessI",termlessI), 
   469        erls = RootEq_erls, srls = Erls, calc = [], errpatts = [],
   470        (*asm_thm = [("sqrt_square_1","")],*)
   471        rules = [Thm  ("real_assoc_1",num_str @{thm real_assoc_1}),
   472                              (* a+(b+c) = a+b+c *)
   473                 Thm  ("real_assoc_2",num_str @{thm real_assoc_2}),
   474                              (* a*(b*c) = a*b*c *)
   475                 Calc ("Groups.plus_class.plus",eval_binop "#add_"),
   476                 Calc ("Groups.minus_class.minus",eval_binop "#sub_"),
   477                 Calc ("Groups.times_class.times",eval_binop "#mult_"),
   478                 Calc ("Fields.inverse_class.divide", eval_cancel "#divide_e"),
   479                 Calc ("NthRoot.sqrt",eval_sqrt "#sqrt_"),
   480                 Calc ("Atools.pow" ,eval_binop "#power_"),
   481                 Thm("real_plus_binom_pow2",num_str @{thm real_plus_binom_pow2}),
   482                 Thm("real_minus_binom_pow2",num_str @{thm real_minus_binom_pow2}),
   483                 Thm("realpow_mul",num_str @{thm realpow_mul}),    
   484                      (* (a * b)^n = a^n * b^n*)
   485                 Thm("sqrt_times_root_1",num_str @{thm sqrt_times_root_1}), 
   486                      (* sqrt b * sqrt c = sqrt(b*c) *)
   487                 Thm("sqrt_times_root_2",num_str @{thm sqrt_times_root_2}),
   488                      (* a * sqrt a * sqrt b = a * sqrt(a*b) *)
   489                 Thm("sqrt_square_2",num_str @{thm sqrt_square_2}),
   490                             (* sqrt (a^^^2) = a *)
   491                 Thm("sqrt_square_1",num_str @{thm sqrt_square_1}) 
   492                             (* sqrt a ^^^ 2 = a *)
   493                 ],
   494        scr = Prog ((term_of o the o (parse thy)) "empty_script")
   495     }:rls);
   496   ruleset' := overwritelthy @{theory} (!ruleset',
   497                           [("rooteq_simplify",rooteq_simplify)
   498                            ]);
   499 *}
   500 setup {* KEStore_Elems.add_rlss
   501   [("rooteq_simplify", (Context.theory_name @{theory}, rooteq_simplify))] *}
   502 ML {*
   503   
   504 (*-------------------------Problem-----------------------*)
   505 (*
   506 (get_pbt ["root'","univariate","equation"]);
   507 show_ptyps(); 
   508 *)
   509 (* ---------root----------- *)
   510 store_pbt
   511  (prep_pbt thy "pbl_equ_univ_root" [] e_pblID
   512  (["root'","univariate","equation"],
   513   [("#Given" ,["equality e_e","solveFor v_v"]),
   514    ("#Where" ,["(lhs e_e) is_rootTerm_in  (v_v::real) | " ^
   515 	       "(rhs e_e) is_rootTerm_in  (v_v::real)"]),
   516    ("#Find"  ,["solutions v_v'i'"]) 
   517   ],
   518   RootEq_prls, SOME "solve (e_e::bool, v_v)",
   519   []));
   520 (* ---------sqrt----------- *)
   521 store_pbt
   522  (prep_pbt thy "pbl_equ_univ_root_sq" [] e_pblID
   523  (["sq","root'","univariate","equation"],
   524   [("#Given" ,["equality e_e","solveFor v_v"]),
   525    ("#Where" ,["( ((lhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
   526                "  ((lhs e_e) is_normSqrtTerm_in (v_v::real))   )  |" ^
   527 	       "( ((rhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
   528                "  ((rhs e_e) is_normSqrtTerm_in (v_v::real))   )"]),
   529    ("#Find"  ,["solutions v_v'i'"]) 
   530   ],
   531   RootEq_prls,  SOME "solve (e_e::bool, v_v)",
   532   [["RootEq","solve_sq_root_equation"]]));
   533 (* ---------normalize----------- *)
   534 store_pbt
   535  (prep_pbt thy "pbl_equ_univ_root_norm" [] e_pblID
   536  (["normalize","root'","univariate","equation"],
   537   [("#Given" ,["equality e_e","solveFor v_v"]),
   538    ("#Where" ,["( ((lhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
   539                "  Not((lhs e_e) is_normSqrtTerm_in (v_v::real)))  | " ^
   540 	       "( ((rhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
   541                "  Not((rhs e_e) is_normSqrtTerm_in (v_v::real)))"]),
   542    ("#Find"  ,["solutions v_v'i'"]) 
   543   ],
   544   RootEq_prls,  SOME "solve (e_e::bool, v_v)",
   545   [["RootEq","norm_sq_root_equation"]]));
   546 
   547 (*-------------------------methods-----------------------*)
   548 (* ---- root 20.8.02 ---*)
   549 store_met
   550  (prep_met thy "met_rooteq" [] e_metID
   551  (["RootEq"],
   552    [],
   553    {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = e_rls, prls=e_rls,
   554     crls=RootEq_crls, errpats = [], nrls = norm_Poly}, "empty_script"));
   555 
   556 (*-- normalize 20.10.02 --*)
   557 store_met
   558  (prep_met thy "met_rooteq_norm" [] e_metID
   559  (["RootEq","norm_sq_root_equation"],
   560    [("#Given" ,["equality e_e","solveFor v_v"]),
   561     ("#Where" ,["( ((lhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
   562                "  Not((lhs e_e) is_normSqrtTerm_in (v_v::real)))  | " ^
   563 	       "( ((rhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
   564                "  Not((rhs e_e) is_normSqrtTerm_in (v_v::real)))"]),
   565     ("#Find"  ,["solutions v_v'i'"])
   566    ],
   567    {rew_ord'="termlessI", rls'=RootEq_erls, srls=e_rls, prls=RootEq_prls,
   568     calc=[], crls=RootEq_crls, errpats = [], nrls = norm_Poly},
   569    "Script Norm_sq_root_equation  (e_e::bool) (v_v::real)  =                " ^
   570     "(let e_e = ((Repeat(Try (Rewrite     makex1_x            False))) @@  " ^
   571     "           (Try (Repeat (Rewrite_Set expand_rootbinoms  False))) @@  " ^ 
   572     "           (Try (Rewrite_Set rooteq_simplify              True)) @@  " ^ 
   573     "           (Try (Repeat (Rewrite_Set make_rooteq        False))) @@  " ^
   574     "           (Try (Rewrite_Set rooteq_simplify              True))) e_e " ^
   575     " in ((SubProblem (RootEq',[univariate,equation],                     " ^
   576     "      [no_met]) [BOOL e_e, REAL v_v])))"
   577    ));
   578 
   579 *}
   580 
   581 ML {*
   582 val -------------------------------------------------- = "00000";
   583 store_met
   584  (prep_met thy "met_rooteq_sq" [] e_metID
   585  (["RootEq","solve_sq_root_equation"],
   586    [("#Given" ,["equality e_e", "solveFor v_v"]),
   587     ("#Where" ,["(((lhs e_e) is_sqrtTerm_in (v_v::real))     & " ^
   588                 " ((lhs e_e) is_normSqrtTerm_in (v_v::real))) |" ^
   589 	        "(((rhs e_e) is_sqrtTerm_in (v_v::real))     & " ^
   590                 " ((rhs e_e) is_normSqrtTerm_in (v_v::real)))"]),
   591     ("#Find"  ,["solutions v_v'i'"])
   592    ],
   593    {rew_ord'="termlessI", rls'=RootEq_erls, srls = rooteq_srls,
   594     prls = RootEq_prls, calc = [], crls=RootEq_crls, errpats = [], nrls = norm_Poly},
   595 "Script Solve_sq_root_equation  (e_e::bool) (v_v::real)  =               " ^
   596 "(let e_e =                                                              " ^
   597 "  ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] sqrt_isolate True)) @@      " ^
   598 "   (Try (Rewrite_Set         rooteq_simplify True))             @@      " ^
   599 "   (Try (Repeat (Rewrite_Set expand_rootbinoms         False))) @@      " ^
   600 "   (Try (Repeat (Rewrite_Set make_rooteq               False))) @@      " ^
   601 "   (Try (Rewrite_Set rooteq_simplify                    True)) ) e_e;   " ^
   602 " (L_L::bool list) =                                                     " ^
   603 "   (if (((lhs e_e) is_sqrtTerm_in v_v) | ((rhs e_e) is_sqrtTerm_in v_v))" ^
   604 "    then (SubProblem (RootEq',[normalize,root',univariate,equation],   " ^
   605 "                      [no_met]) [BOOL e_e, REAL v_v])                   " ^
   606 "    else (SubProblem (RootEq',[univariate,equation], [no_met])         " ^
   607 "                     [BOOL e_e, REAL v_v]))                             " ^
   608 "in Check_elementwise L_L {(v_v::real). Assumptions})"
   609  ));
   610 *}
   611 
   612 ML {*
   613 (*-- right 28.08.02 --*)
   614 store_met
   615  (prep_met thy "met_rooteq_sq_right" [] e_metID
   616  (["RootEq","solve_right_sq_root_equation"],
   617    [("#Given" ,["equality e_e","solveFor v_v"]),
   618     ("#Where" ,["(rhs e_e) is_sqrtTerm_in v_v"]),
   619     ("#Find"  ,["solutions v_v'i'"])
   620    ],
   621    {rew_ord' = "termlessI", rls' = RootEq_erls, srls = e_rls, 
   622     prls = RootEq_prls, calc = [], crls = RootEq_crls, errpats = [], nrls = norm_Poly},
   623   "Script Solve_right_sq_root_equation  (e_e::bool) (v_v::real)  =           " ^
   624    "(let e_e =                                                               " ^
   625    "    ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] r_sqrt_isolate  False)) @@ " ^
   626    "    (Try (Rewrite_Set                       rooteq_simplify False)) @@   " ^
   627    "    (Try (Repeat (Rewrite_Set expand_rootbinoms            False))) @@   " ^
   628    "    (Try (Repeat (Rewrite_Set make_rooteq                  False))) @@   " ^
   629    "    (Try (Rewrite_Set rooteq_simplify                       False))) e_e " ^
   630    " in if ((rhs e_e) is_sqrtTerm_in v_v)                                    " ^
   631    " then (SubProblem (RootEq',[normalize,root',univariate,equation],       " ^
   632    "       [no_met]) [BOOL e_e, REAL v_v])                                   " ^
   633    " else ((SubProblem (RootEq',[univariate,equation],                      " ^
   634    "        [no_met]) [BOOL e_e, REAL v_v])))"
   635  ));
   636 val --------------------------------------------------+++ = "33333";
   637 
   638 (*-- left 28.08.02 --*)
   639 store_met
   640  (prep_met thy "met_rooteq_sq_left" [] e_metID
   641  (["RootEq","solve_left_sq_root_equation"],
   642    [("#Given" ,["equality e_e","solveFor v_v"]),
   643     ("#Where" ,["(lhs e_e) is_sqrtTerm_in v_v"]),
   644     ("#Find"  ,["solutions v_v'i'"])
   645    ],
   646    {rew_ord'="termlessI",
   647     rls'=RootEq_erls,
   648     srls=e_rls,
   649     prls=RootEq_prls,
   650     calc=[],
   651     crls=RootEq_crls, errpats = [], nrls = norm_Poly},
   652     "Script Solve_left_sq_root_equation  (e_e::bool) (v_v::real)  =          " ^
   653     "(let e_e =                                                             " ^
   654     "  ((Try (Rewrite_Set_Inst [(bdv,v_v::real)] l_sqrt_isolate  False)) @@ " ^
   655     "  (Try (Rewrite_Set                       rooteq_simplify False)) @@  " ^
   656     "  (Try (Repeat (Rewrite_Set expand_rootbinoms            False))) @@  " ^
   657     "  (Try (Repeat (Rewrite_Set make_rooteq                  False))) @@  " ^
   658     "  (Try (Rewrite_Set rooteq_simplify                       False))) e_e " ^
   659     " in if ((lhs e_e) is_sqrtTerm_in v_v)                                   " ^ 
   660     " then (SubProblem (RootEq',[normalize,root',univariate,equation],      " ^
   661     "       [no_met]) [BOOL e_e, REAL v_v])                                " ^
   662     " else ((SubProblem (RootEq',[univariate,equation],                    " ^
   663     "        [no_met]) [BOOL e_e, REAL v_v])))"
   664    ));
   665 val --------------------------------------------------++++ = "44444";
   666 
   667 calclist':= overwritel (!calclist', 
   668    [("is_rootTerm_in", ("RootEq.is'_rootTerm'_in", 
   669 			eval_is_rootTerm_in"")),
   670     ("is_sqrtTerm_in", ("RootEq.is'_sqrtTerm'_in", 
   671 			eval_is_sqrtTerm_in"")),
   672     ("is_normSqrtTerm_in", ("RootEq.is_normSqrtTerm_in", 
   673 				 eval_is_normSqrtTerm_in""))
   674     ]);(*("", ("", "")),*)
   675 *}
   676 
   677 end