src/Tools/isac/Knowledge/RootEq.thy
author wenzelm
Wed, 26 May 2021 16:24:05 +0200
changeset 60286 31efa1b39a20
parent 60278 343efa173023
child 60289 a7b88fc19a92
permissions -rw-r--r--
command 'setup_rule' semantic equivalent for KEStore_Elems.add_rlss;
neuper@37906
     1
(*.(c) by Richard Lang, 2003 .*)
neuper@37906
     2
(* collecting all knowledge for Root Equations
neuper@37906
     3
   created by: rlang 
neuper@37906
     4
         date: 02.08
neuper@37906
     5
   changed by: rlang
neuper@37906
     6
   last change by: rlang
neuper@37906
     7
             date: 02.11.14
neuper@37906
     8
*)
neuper@37906
     9
s1210629013@55339
    10
theory RootEq imports Root Equation begin
neuper@37906
    11
wneuper@59472
    12
text \<open>univariate equations containing real square roots:
neuper@42398
    13
  This type of equations has been used to bootstrap Lucas-Interpretation.
neuper@42398
    14
  In 2003 this type has been extended and integrated into ISAC's equation solver
neuper@42398
    15
  by Richard Lang in 2003.
neuper@42398
    16
  The assumptions (all containing "<") didn't pass the xml-parsers at the 
neuper@42398
    17
  interface between math-engine and front-end.
neuper@42398
    18
  The migration Isabelle2002 --> 2011 dropped this type of equation, see
neuper@42398
    19
  test/../rooteq.sml, rootrateq.sml.
wneuper@59472
    20
\<close>
neuper@42398
    21
wneuper@59526
    22
subsection \<open>consts definition for predicates\<close>
neuper@37906
    23
consts
walther@60278
    24
  is_rootTerm_in     :: "[real, real] => bool" ("_ is'_rootTerm'_in _")
walther@60278
    25
  is_sqrtTerm_in     :: "[real, real] => bool" ("_ is'_sqrtTerm'_in _") 
walther@60278
    26
  is_normSqrtTerm_in :: "[real, real] => bool" ("_ is'_normSqrtTerm'_in _") 
neuper@37906
    27
 
wneuper@59526
    28
subsection \<open>theorems not yet adopted from Isabelle\<close>
neuper@52148
    29
axiomatization where
wneuper@59370
    30
(* normalise *)
walther@60242
    31
  makex1_x:            "a \<up> 1  = a"   and
neuper@52148
    32
  real_assoc_1:        "a+(b+c) = a+b+c" and
neuper@52148
    33
  real_assoc_2:        "a*(b*c) = a*b*c" and
neuper@37906
    34
neuper@37906
    35
  (* simplification of root*)
walther@60242
    36
  sqrt_square_1:       "[|0 <= a|] ==>  (sqrt a) \<up> 2 = a" and
walther@60242
    37
  sqrt_square_2:       "sqrt (a  \<up>  2) = a" and
neuper@52148
    38
  sqrt_times_root_1:   "sqrt a * sqrt b = sqrt(a*b)" and
neuper@52148
    39
  sqrt_times_root_2:   "a * sqrt b * sqrt c = a * sqrt(b*c)" and
neuper@37906
    40
neuper@37906
    41
  (* isolate one root on the LEFT or RIGHT hand side of the equation *)
neuper@37983
    42
  sqrt_isolate_l_add1: "[|bdv occurs_in c|] ==> 
neuper@52148
    43
   (a + b*sqrt(c) = d) = (b * sqrt(c) = d+ (-1) * a)" and
neuper@37983
    44
  sqrt_isolate_l_add2: "[|bdv occurs_in c|] ==>
neuper@52148
    45
   (a + sqrt(c) = d) = ((sqrt(c) = d+ (-1) * a))" and
neuper@37983
    46
  sqrt_isolate_l_add3: "[|bdv occurs_in c|] ==>
neuper@52148
    47
   (a + b*(e/sqrt(c)) = d) = (b * (e/sqrt(c)) = d + (-1) * a)" and
neuper@37983
    48
  sqrt_isolate_l_add4: "[|bdv occurs_in c|] ==>
neuper@52148
    49
   (a + b/(f*sqrt(c)) = d) = (b / (f*sqrt(c)) = d + (-1) * a)" and
neuper@37983
    50
  sqrt_isolate_l_add5: "[|bdv occurs_in c|] ==>
neuper@52148
    51
   (a + b*(e/(f*sqrt(c))) = d) = (b * (e/(f*sqrt(c))) = d+ (-1) * a)" and
neuper@37983
    52
  sqrt_isolate_l_add6: "[|bdv occurs_in c|] ==>
neuper@52148
    53
   (a + b/sqrt(c) = d) = (b / sqrt(c) = d+ (-1) * a)" and
neuper@37983
    54
  sqrt_isolate_r_add1: "[|bdv occurs_in f|] ==>
neuper@52148
    55
   (a = d + e*sqrt(f)) = (a + (-1) * d = e*sqrt(f))" and
neuper@37983
    56
  sqrt_isolate_r_add2: "[|bdv occurs_in f|] ==>
neuper@52148
    57
   (a = d + sqrt(f)) = (a + (-1) * d = sqrt(f))" and
wneuper@59370
    58
 (* small hack: thm 3,5,6 are not needed if rootnormalise is well done*)
neuper@37983
    59
  sqrt_isolate_r_add3: "[|bdv occurs_in f|] ==>
neuper@52148
    60
   (a = d + e*(g/sqrt(f))) = (a + (-1) * d = e*(g/sqrt(f)))" and
neuper@37983
    61
  sqrt_isolate_r_add4: "[|bdv occurs_in f|] ==>
neuper@52148
    62
   (a = d + g/sqrt(f)) = (a + (-1) * d = g/sqrt(f))" and
neuper@37983
    63
  sqrt_isolate_r_add5: "[|bdv occurs_in f|] ==>
neuper@52148
    64
   (a = d + e*(g/(h*sqrt(f)))) = (a + (-1) * d = e*(g/(h*sqrt(f))))" and
neuper@37983
    65
  sqrt_isolate_r_add6: "[|bdv occurs_in f|] ==>
neuper@52148
    66
   (a = d + g/(h*sqrt(f))) = (a + (-1) * d = g/(h*sqrt(f)))" and
neuper@37906
    67
 
neuper@37906
    68
  (* eliminate isolates sqrt *)
neuper@37983
    69
  sqrt_square_equation_both_1: "[|bdv occurs_in b; bdv occurs_in d|] ==> 
neuper@37950
    70
   ( (sqrt a + sqrt b         = sqrt c + sqrt d) = 
neuper@52148
    71
     (a+2*sqrt(a)*sqrt(b)+b  = c+2*sqrt(c)*sqrt(d)+d))" and
neuper@37983
    72
  sqrt_square_equation_both_2: "[|bdv occurs_in b; bdv occurs_in d|] ==> 
neuper@37950
    73
   ( (sqrt a - sqrt b           = sqrt c + sqrt d) = 
neuper@52148
    74
     (a - 2*sqrt(a)*sqrt(b)+b  = c+2*sqrt(c)*sqrt(d)+d))" and
neuper@37983
    75
  sqrt_square_equation_both_3: "[|bdv occurs_in b; bdv occurs_in d|] ==> 
neuper@37950
    76
   ( (sqrt a + sqrt b           = sqrt c - sqrt d) = 
neuper@52148
    77
     (a + 2*sqrt(a)*sqrt(b)+b  = c - 2*sqrt(c)*sqrt(d)+d))" and
neuper@37983
    78
  sqrt_square_equation_both_4: "[|bdv occurs_in b; bdv occurs_in d|] ==> 
neuper@37950
    79
   ( (sqrt a - sqrt b           = sqrt c - sqrt d) = 
neuper@52148
    80
     (a - 2*sqrt(a)*sqrt(b)+b  = c - 2*sqrt(c)*sqrt(d)+d))" and
neuper@37983
    81
  sqrt_square_equation_left_1: "[|bdv occurs_in a; 0 <= a; 0 <= b|] ==>
walther@60242
    82
   ( (sqrt (a) = b) = (a = (b \<up> 2)))" and
neuper@37983
    83
  sqrt_square_equation_left_2: "[|bdv occurs_in a; 0 <= a; 0 <= b*c|] ==> 
walther@60242
    84
   ( (c*sqrt(a) = b) = (c \<up> 2*a = b \<up> 2))" and
neuper@37983
    85
  sqrt_square_equation_left_3: "[|bdv occurs_in a; 0 <= a; 0 <= b*c|] ==> 
walther@60242
    86
   ( c/sqrt(a) = b) = (c \<up> 2 / a = b \<up> 2)" and
wneuper@59370
    87
  (* small hack: thm 4-6 are not needed if rootnormalise is well done*)
neuper@37983
    88
  sqrt_square_equation_left_4: "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d|] ==> 
walther@60242
    89
   ( (c*(d/sqrt (a)) = b) = (c \<up> 2*(d \<up> 2/a) = b \<up> 2))" and
neuper@37983
    90
  sqrt_square_equation_left_5: "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d|] ==> 
walther@60242
    91
   ( c/(d*sqrt(a)) = b) = (c \<up> 2 / (d \<up> 2*a) = b \<up> 2)" and
neuper@37983
    92
  sqrt_square_equation_left_6: "[|bdv occurs_in a; 0 <= a; 0 <= b*c*d*e|] ==> 
walther@60242
    93
   ( (c*(d/(e*sqrt (a))) = b) = (c \<up> 2*(d \<up> 2/(e \<up> 2*a)) = b \<up> 2))" and
neuper@37983
    94
  sqrt_square_equation_right_1:  "[|bdv occurs_in b; 0 <= a; 0 <= b|] ==> 
walther@60242
    95
   ( (a = sqrt (b)) = (a \<up> 2 = b))" and
neuper@37983
    96
  sqrt_square_equation_right_2: "[|bdv occurs_in b; 0 <= a*c; 0 <= b|] ==> 
walther@60242
    97
   ( (a = c*sqrt (b)) = ((a \<up> 2) = c \<up> 2*b))" and
neuper@37983
    98
  sqrt_square_equation_right_3: "[|bdv occurs_in b; 0 <= a*c; 0 <= b|] ==> 
walther@60242
    99
   ( (a = c/sqrt (b)) = (a \<up> 2 = c \<up> 2/b))" and
wneuper@59370
   100
 (* small hack: thm 4-6 are not needed if rootnormalise is well done*)
neuper@37983
   101
  sqrt_square_equation_right_4: "[|bdv occurs_in b; 0 <= a*c*d; 0 <= b|] ==> 
walther@60242
   102
   ( (a = c*(d/sqrt (b))) = ((a \<up> 2) = c \<up> 2*(d \<up> 2/b)))" and
neuper@37983
   103
  sqrt_square_equation_right_5: "[|bdv occurs_in b; 0 <= a*c*d; 0 <= b|] ==> 
walther@60242
   104
   ( (a = c/(d*sqrt (b))) = (a \<up> 2 = c \<up> 2/(d \<up> 2*b)))" and
neuper@37983
   105
  sqrt_square_equation_right_6: "[|bdv occurs_in b; 0 <= a*c*d*e; 0 <= b|] ==> 
walther@60242
   106
   ( (a = c*(d/(e*sqrt (b)))) = ((a \<up> 2) = c \<up> 2*(d \<up> 2/(e \<up> 2*b))))"
neuper@37950
   107
wneuper@59526
   108
subsection \<open>predicates\<close>
wneuper@59472
   109
ML \<open>
neuper@37972
   110
val thy = @{theory};
neuper@37972
   111
neuper@37950
   112
(* true if bdv is under sqrt of a Equation*)
neuper@37950
   113
fun is_rootTerm_in t v = 
wneuper@59525
   114
  let 
wneuper@59526
   115
   	fun findroot (t as (_ $ _ $ _ $ _)) _ = raise TERM ( "is_rootTerm_in", [t])
neuper@37950
   116
	  (* at the moment there is no term like this, but ....*)
walther@59603
   117
	  | findroot (Const ("Root.nroot",_) $ _ $ t3) v = Prog_Expr.occurs_in v t3
wneuper@59526
   118
	  | findroot (_ $ t2 $ t3) v = findroot t2 v orelse findroot t3 v
walther@59603
   119
	  | findroot (Const ("NthRoot.sqrt", _) $ t2) v = Prog_Expr.occurs_in v t2
wneuper@59526
   120
	  | findroot (_ $ t2) v = findroot t2 v
neuper@37950
   121
	  | findroot _ _ = false;
wneuper@59526
   122
  in
wneuper@59526
   123
	  findroot t v
wneuper@59526
   124
  end;
neuper@37950
   125
wneuper@59526
   126
fun is_sqrtTerm_in t v = 
wneuper@59526
   127
  let 
wneuper@59526
   128
    fun findsqrt (t as (_ $ _ $ _ $ _)) _ = raise TERM ("is_sqrteqation_in", [t])
wneuper@59526
   129
    (* at the moment there is no term like this, but ....*)
wneuper@59526
   130
    | findsqrt (_ $ t1 $ t2) v = (findsqrt t1 v) orelse (findsqrt t2 v)
walther@59603
   131
    | findsqrt (Const ("NthRoot.sqrt", _) $ a) v = Prog_Expr.occurs_in v a
wneuper@59526
   132
    | findsqrt (_ $ t1) v = findsqrt t1 v
wneuper@59526
   133
    | findsqrt _ _ = false;
wneuper@59526
   134
  in
wneuper@59526
   135
   findsqrt t v
wneuper@59526
   136
  end;
neuper@37950
   137
neuper@37950
   138
(* RL: 030518: Is in the rightest subterm of a term a sqrt with bdv,
wneuper@59370
   139
and the subterm ist connected with + or * --> is normalised*)
wneuper@59526
   140
fun is_normSqrtTerm_in t v =
wneuper@59526
   141
  let
wneuper@59526
   142
     fun isnorm (t as (_ $ _ $ _ $ _)) _ = raise TERM ("is_normSqrtTerm_in", [t])
wneuper@59526
   143
         (* at the moment there is no term like this, but ....*)
wneuper@59526
   144
       | isnorm (Const ("Groups.plus_class.plus",_) $ _ $ t2) v = is_sqrtTerm_in t2 v
wneuper@59526
   145
       | isnorm (Const ("Groups.times_class.times",_) $ _ $ t2) v = is_sqrtTerm_in t2 v
wneuper@59526
   146
       | isnorm (Const ("Groups.minus_class.minus",_) $ _ $ _) _ = false
wneuper@59526
   147
       | isnorm (Const ("Rings.divide_class.divide",_) $ t1 $ t2) v =
wneuper@59526
   148
         is_sqrtTerm_in t1 v orelse is_sqrtTerm_in t2 v
walther@59603
   149
       | isnorm (Const ("NthRoot.sqrt", _) $ t1) v = Prog_Expr.occurs_in v t1
wneuper@59526
   150
	     | isnorm (_ $ t1) v = is_sqrtTerm_in t1 v
wneuper@59526
   151
       | isnorm _ _ = false;
wneuper@59526
   152
   in
wneuper@59526
   153
     isnorm t v
wneuper@59526
   154
   end;
neuper@37950
   155
walther@60278
   156
fun eval_is_rootTerm_in _ _ (p as (Const ("RootEq.is_rootTerm_in",_) $ t $ v)) _ =
wneuper@59526
   157
    if is_rootTerm_in t v
walther@59868
   158
    then SOME ((UnparseC.term p) ^ " = True", HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   159
    else SOME ((UnparseC.term p) ^ " = True", HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
neuper@38015
   160
  | eval_is_rootTerm_in _ _ _ _ = ((*tracing"### nichts matcht";*) NONE);
neuper@37950
   161
walther@60278
   162
fun eval_is_sqrtTerm_in _ _ (p as (Const ("RootEq.is_sqrtTerm_in",_) $ t $ v)) _ =
wneuper@59526
   163
    if is_sqrtTerm_in t v
walther@59868
   164
    then SOME ((UnparseC.term p) ^ " = True", HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   165
    else SOME ((UnparseC.term p) ^ " = True", HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
neuper@38015
   166
  | eval_is_sqrtTerm_in _ _ _ _ = ((*tracing"### nichts matcht";*) NONE);
neuper@37950
   167
walther@60278
   168
fun eval_is_normSqrtTerm_in _ _ (p as (Const ("RootEq.is_normSqrtTerm_in",_) $ t $ v)) _ =
wneuper@59526
   169
    if is_normSqrtTerm_in t v
walther@59868
   170
    then SOME ((UnparseC.term p) ^ " = True", HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   171
    else SOME ((UnparseC.term p) ^ " = True", HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
neuper@38015
   172
  | eval_is_normSqrtTerm_in _ _ _ _ = ((*tracing"### nichts matcht";*) NONE);
wneuper@59526
   173
\<close>
wneuper@59526
   174
setup \<open>KEStore_Elems.add_calcs
walther@60278
   175
  [("is_rootTerm_in", ("RootEq.is_rootTerm_in", eval_is_rootTerm_in"")),
walther@60278
   176
    ("is_sqrtTerm_in", ("RootEq.is_sqrtTerm_in", eval_is_sqrtTerm_in"")),
wneuper@59526
   177
    ("is_normSqrtTerm_in", ("RootEq.is_normSqrtTerm_in", eval_is_normSqrtTerm_in""))]\<close>
neuper@37950
   178
wneuper@59526
   179
subsection \<open>rule-sets\<close>
wneuper@59526
   180
ML \<open>
neuper@37950
   181
val RootEq_prls =(*15.10.02:just the following order due to subterm evaluation*)
walther@59852
   182
  Rule_Set.append_rules "RootEq_prls" Rule_Set.empty 
walther@59878
   183
	     [Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),
walther@59878
   184
	      Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches ""),
walther@59878
   185
	      Rule.Eval ("Prog_Expr.lhs"    , Prog_Expr.eval_lhs ""),
walther@59878
   186
	      Rule.Eval ("Prog_Expr.rhs"    , Prog_Expr.eval_rhs ""),
walther@60278
   187
	      Rule.Eval ("RootEq.is_sqrtTerm_in", eval_is_sqrtTerm_in ""),
walther@60278
   188
	      Rule.Eval ("RootEq.is_rootTerm_in", eval_is_rootTerm_in ""),
walther@60278
   189
	      Rule.Eval ("RootEq.is_normSqrtTerm_in", eval_is_normSqrtTerm_in ""),
walther@59878
   190
	      Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
walther@59871
   191
	      Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
walther@59871
   192
	      Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false}),
walther@59871
   193
	      Rule.Thm ("and_true", ThmC.numerals_to_Free @{thm and_true}),
walther@59871
   194
	      Rule.Thm ("and_false", ThmC.numerals_to_Free @{thm and_false}),
walther@59871
   195
	      Rule.Thm ("or_true", ThmC.numerals_to_Free @{thm or_true}),
walther@59871
   196
	      Rule.Thm ("or_false", ThmC.numerals_to_Free @{thm or_false})
neuper@37950
   197
	      ];
neuper@37950
   198
neuper@37950
   199
val RootEq_erls =
walther@59852
   200
  Rule_Set.append_rules "RootEq_erls" Root_erls
walther@59871
   201
    [Rule.Thm ("divide_divide_eq_left", ThmC.numerals_to_Free @{thm divide_divide_eq_left})];
neuper@37950
   202
neuper@37950
   203
val RootEq_crls = 
walther@59852
   204
  Rule_Set.append_rules "RootEq_crls" Root_crls
walther@59871
   205
    [Rule.Thm ("divide_divide_eq_left", ThmC.numerals_to_Free @{thm divide_divide_eq_left})];
neuper@37950
   206
neuper@37950
   207
val rooteq_srls = 
walther@59852
   208
  Rule_Set.append_rules "rooteq_srls" Rule_Set.empty
walther@60278
   209
    [Rule.Eval ("RootEq.is_sqrtTerm_in", eval_is_sqrtTerm_in ""),
walther@60278
   210
     Rule.Eval ("RootEq.is_normSqrtTerm_in", eval_is_normSqrtTerm_in""),
walther@60278
   211
     Rule.Eval ("RootEq.is_rootTerm_in", eval_is_rootTerm_in "")];
wneuper@59472
   212
\<close>
wneuper@59472
   213
ML \<open>
neuper@37950
   214
neuper@37950
   215
(*isolate the bound variable in an sqrt equation; 'bdv' is a meta-constant*)
s1210629013@55444
   216
 val sqrt_isolate = prep_rls'(
walther@59851
   217
  Rule_Def.Repeat {id = "sqrt_isolate", preconds = [], rew_ord = ("termlessI",termlessI), 
walther@59851
   218
       erls = RootEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
neuper@37950
   219
       rules = [
walther@59871
   220
       Rule.Thm("sqrt_square_1",ThmC.numerals_to_Free @{thm sqrt_square_1}),
walther@60260
   221
                     (* (sqrt a) \<up> 2 -> a *)
walther@59871
   222
       Rule.Thm("sqrt_square_2",ThmC.numerals_to_Free @{thm sqrt_square_2}),
walther@60260
   223
                     (* sqrt (a \<up> 2) -> a *)
walther@59871
   224
       Rule.Thm("sqrt_times_root_1",ThmC.numerals_to_Free @{thm sqrt_times_root_1}),
neuper@37950
   225
            (* sqrt a sqrt b -> sqrt(ab) *)
walther@59871
   226
       Rule.Thm("sqrt_times_root_2",ThmC.numerals_to_Free @{thm sqrt_times_root_2}),
neuper@37950
   227
            (* a sqrt b sqrt c -> a sqrt(bc) *)
wneuper@59416
   228
       Rule.Thm("sqrt_square_equation_both_1",
walther@59871
   229
           ThmC.numerals_to_Free @{thm sqrt_square_equation_both_1}),
neuper@37950
   230
       (* (sqrt a + sqrt b  = sqrt c + sqrt d) -> 
neuper@37950
   231
            (a+2*sqrt(a)*sqrt(b)+b) = c+2*sqrt(c)*sqrt(d)+d) *)
wneuper@59416
   232
       Rule.Thm("sqrt_square_equation_both_2",
walther@59871
   233
            ThmC.numerals_to_Free @{thm sqrt_square_equation_both_2}),
neuper@37950
   234
       (* (sqrt a - sqrt b  = sqrt c + sqrt d) -> 
neuper@37950
   235
            (a-2*sqrt(a)*sqrt(b)+b) = c+2*sqrt(c)*sqrt(d)+d) *)
wneuper@59416
   236
       Rule.Thm("sqrt_square_equation_both_3",
walther@59871
   237
            ThmC.numerals_to_Free @{thm sqrt_square_equation_both_3}),
neuper@37950
   238
       (* (sqrt a + sqrt b  = sqrt c - sqrt d) -> 
neuper@37950
   239
            (a+2*sqrt(a)*sqrt(b)+b) = c-2*sqrt(c)*sqrt(d)+d) *)
wneuper@59416
   240
       Rule.Thm("sqrt_square_equation_both_4",
walther@59871
   241
            ThmC.numerals_to_Free @{thm sqrt_square_equation_both_4}),
neuper@37950
   242
       (* (sqrt a - sqrt b  = sqrt c - sqrt d) -> 
neuper@37950
   243
            (a-2*sqrt(a)*sqrt(b)+b) = c-2*sqrt(c)*sqrt(d)+d) *)
wneuper@59416
   244
       Rule.Thm("sqrt_isolate_l_add1",
walther@59871
   245
            ThmC.numerals_to_Free @{thm sqrt_isolate_l_add1}), 
neuper@37950
   246
       (* a+b*sqrt(x)=d -> b*sqrt(x) = d-a *)
wneuper@59416
   247
       Rule.Thm("sqrt_isolate_l_add2",
walther@59871
   248
            ThmC.numerals_to_Free @{thm sqrt_isolate_l_add2}), 
neuper@37950
   249
       (* a+  sqrt(x)=d ->   sqrt(x) = d-a *)
wneuper@59416
   250
       Rule.Thm("sqrt_isolate_l_add3",
walther@59871
   251
            ThmC.numerals_to_Free @{thm sqrt_isolate_l_add3}), 
neuper@37950
   252
       (* a+b*c/sqrt(x)=d->b*c/sqrt(x)=d-a *)
wneuper@59416
   253
       Rule.Thm("sqrt_isolate_l_add4",
walther@59871
   254
            ThmC.numerals_to_Free @{thm sqrt_isolate_l_add4}), 
neuper@37950
   255
       (* a+c/sqrt(x)=d -> c/sqrt(x) = d-a *)
wneuper@59416
   256
       Rule.Thm("sqrt_isolate_l_add5",
walther@59871
   257
            ThmC.numerals_to_Free @{thm sqrt_isolate_l_add5}), 
neuper@37950
   258
       (* a+b*c/f*sqrt(x)=d->b*c/f*sqrt(x)=d-a *)
wneuper@59416
   259
       Rule.Thm("sqrt_isolate_l_add6",
walther@59871
   260
            ThmC.numerals_to_Free @{thm sqrt_isolate_l_add6}), 
neuper@37950
   261
       (* a+c/f*sqrt(x)=d -> c/f*sqrt(x) = d-a *)
walther@59871
   262
       (*Rule.Thm("sqrt_isolate_l_div",ThmC.numerals_to_Free @{thm sqrt_isolate_l_div}),*)
neuper@37950
   263
         (* b*sqrt(x) = d sqrt(x) d/b *)
wneuper@59416
   264
       Rule.Thm("sqrt_isolate_r_add1",
walther@59871
   265
            ThmC.numerals_to_Free @{thm sqrt_isolate_r_add1}),
neuper@37950
   266
       (* a= d+e*sqrt(x) -> a-d=e*sqrt(x) *)
wneuper@59416
   267
       Rule.Thm("sqrt_isolate_r_add2",
walther@59871
   268
            ThmC.numerals_to_Free @{thm sqrt_isolate_r_add2}),
neuper@37950
   269
       (* a= d+  sqrt(x) -> a-d=  sqrt(x) *)
wneuper@59416
   270
       Rule.Thm("sqrt_isolate_r_add3",
walther@59871
   271
            ThmC.numerals_to_Free @{thm sqrt_isolate_r_add3}),
neuper@37950
   272
       (* a=d+e*g/sqrt(x)->a-d=e*g/sqrt(x)*)
wneuper@59416
   273
       Rule.Thm("sqrt_isolate_r_add4",
walther@59871
   274
            ThmC.numerals_to_Free @{thm sqrt_isolate_r_add4}),
neuper@37950
   275
       (* a= d+g/sqrt(x) -> a-d=g/sqrt(x) *)
wneuper@59416
   276
       Rule.Thm("sqrt_isolate_r_add5",
walther@59871
   277
            ThmC.numerals_to_Free @{thm sqrt_isolate_r_add5}),
neuper@37950
   278
       (* a=d+e*g/h*sqrt(x)->a-d=e*g/h*sqrt(x)*)
wneuper@59416
   279
       Rule.Thm("sqrt_isolate_r_add6",
walther@59871
   280
            ThmC.numerals_to_Free @{thm sqrt_isolate_r_add6}),
neuper@37950
   281
       (* a= d+g/h*sqrt(x) -> a-d=g/h*sqrt(x) *)
walther@59871
   282
       (*Rule.Thm("sqrt_isolate_r_div",ThmC.numerals_to_Free @{thm sqrt_isolate_r_div}),*)
neuper@37950
   283
         (* a=e*sqrt(x) -> a/e = sqrt(x) *)
wneuper@59416
   284
       Rule.Thm("sqrt_square_equation_left_1",
walther@59871
   285
            ThmC.numerals_to_Free @{thm sqrt_square_equation_left_1}),   
neuper@37950
   286
       (* sqrt(x)=b -> x=b^2 *)
wneuper@59416
   287
       Rule.Thm("sqrt_square_equation_left_2",
walther@59871
   288
            ThmC.numerals_to_Free @{thm sqrt_square_equation_left_2}),   
neuper@37950
   289
       (* c*sqrt(x)=b -> c^2*x=b^2 *)
walther@59871
   290
       Rule.Thm("sqrt_square_equation_left_3",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_3}),  
neuper@37950
   291
	      (* c/sqrt(x)=b -> c^2/x=b^2 *)
walther@59871
   292
       Rule.Thm("sqrt_square_equation_left_4",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_4}),
neuper@37950
   293
	      (* c*d/sqrt(x)=b -> c^2*d^2/x=b^2 *)
walther@59871
   294
       Rule.Thm("sqrt_square_equation_left_5",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_5}),
neuper@37950
   295
	      (* c/d*sqrt(x)=b -> c^2/d^2x=b^2 *)
walther@59871
   296
       Rule.Thm("sqrt_square_equation_left_6",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_6}),
neuper@37950
   297
	      (* c*d/g*sqrt(x)=b -> c^2*d^2/g^2x=b^2 *)
walther@59871
   298
       Rule.Thm("sqrt_square_equation_right_1",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_1}),
neuper@37950
   299
	      (* a=sqrt(x) ->a^2=x *)
walther@59871
   300
       Rule.Thm("sqrt_square_equation_right_2",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_2}),
neuper@37950
   301
	      (* a=c*sqrt(x) ->a^2=c^2*x *)
walther@59871
   302
       Rule.Thm("sqrt_square_equation_right_3",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_3}),
neuper@37950
   303
	      (* a=c/sqrt(x) ->a^2=c^2/x *)
walther@59871
   304
       Rule.Thm("sqrt_square_equation_right_4",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_4}),
neuper@37950
   305
	      (* a=c*d/sqrt(x) ->a^2=c^2*d^2/x *)
walther@59871
   306
       Rule.Thm("sqrt_square_equation_right_5",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_5}),
neuper@37950
   307
	      (* a=c/e*sqrt(x) ->a^2=c^2/e^2x *)
walther@59871
   308
       Rule.Thm("sqrt_square_equation_right_6",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_6})
neuper@37950
   309
	      (* a=c*d/g*sqrt(x) ->a^2=c^2*d^2/g^2*x *)
walther@59878
   310
       ],scr = Rule.Empty_Prog
wneuper@59406
   311
      });
neuper@52125
   312
wneuper@59526
   313
\<close> ML \<open>
neuper@37950
   314
(*isolate the bound variable in an sqrt left equation; 'bdv' is a meta-constant*)
s1210629013@55444
   315
 val l_sqrt_isolate = prep_rls'(
walther@59851
   316
     Rule_Def.Repeat {id = "l_sqrt_isolate", preconds = [], 
neuper@37950
   317
	  rew_ord = ("termlessI",termlessI), 
walther@59851
   318
          erls = RootEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
neuper@37950
   319
     rules = [
walther@59871
   320
     Rule.Thm("sqrt_square_1",ThmC.numerals_to_Free @{thm sqrt_square_1}),
walther@60260
   321
                            (* (sqrt a) \<up> 2 -> a *)
walther@59871
   322
     Rule.Thm("sqrt_square_2",ThmC.numerals_to_Free @{thm sqrt_square_2}),
walther@60260
   323
                            (* sqrt (a \<up> 2) -> a *)
walther@59871
   324
     Rule.Thm("sqrt_times_root_1",ThmC.numerals_to_Free @{thm sqrt_times_root_1}),
neuper@37950
   325
            (* sqrt a sqrt b -> sqrt(ab) *)
walther@59871
   326
     Rule.Thm("sqrt_times_root_2",ThmC.numerals_to_Free @{thm sqrt_times_root_2}),
neuper@37950
   327
        (* a sqrt b sqrt c -> a sqrt(bc) *)
walther@59871
   328
     Rule.Thm("sqrt_isolate_l_add1",ThmC.numerals_to_Free @{thm sqrt_isolate_l_add1}),
neuper@37950
   329
        (* a+b*sqrt(x)=d -> b*sqrt(x) = d-a *)
walther@59871
   330
     Rule.Thm("sqrt_isolate_l_add2",ThmC.numerals_to_Free @{thm sqrt_isolate_l_add2}),
neuper@37950
   331
        (* a+  sqrt(x)=d ->   sqrt(x) = d-a *)
walther@59871
   332
     Rule.Thm("sqrt_isolate_l_add3",ThmC.numerals_to_Free @{thm sqrt_isolate_l_add3}),
neuper@37950
   333
        (* a+b*c/sqrt(x)=d->b*c/sqrt(x)=d-a *)
walther@59871
   334
     Rule.Thm("sqrt_isolate_l_add4",ThmC.numerals_to_Free @{thm sqrt_isolate_l_add4}),
neuper@37950
   335
        (* a+c/sqrt(x)=d -> c/sqrt(x) = d-a *)
walther@59871
   336
     Rule.Thm("sqrt_isolate_l_add5",ThmC.numerals_to_Free @{thm sqrt_isolate_l_add5}),
neuper@37950
   337
        (* a+b*c/f*sqrt(x)=d->b*c/f*sqrt(x)=d-a *)
walther@59871
   338
     Rule.Thm("sqrt_isolate_l_add6",ThmC.numerals_to_Free @{thm sqrt_isolate_l_add6}),
neuper@37950
   339
        (* a+c/f*sqrt(x)=d -> c/f*sqrt(x) = d-a *)
walther@59871
   340
   (*Rule.Thm("sqrt_isolate_l_div",ThmC.numerals_to_Free @{thm sqrt_isolate_l_div}),*)
neuper@37950
   341
        (* b*sqrt(x) = d sqrt(x) d/b *)
walther@59871
   342
     Rule.Thm("sqrt_square_equation_left_1",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_1}),
neuper@37950
   343
	      (* sqrt(x)=b -> x=b^2 *)
walther@59871
   344
     Rule.Thm("sqrt_square_equation_left_2",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_2}),
neuper@37950
   345
	      (* a*sqrt(x)=b -> a^2*x=b^2*)
walther@59871
   346
     Rule.Thm("sqrt_square_equation_left_3",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_3}),   
neuper@37950
   347
	      (* c/sqrt(x)=b -> c^2/x=b^2 *)
walther@59871
   348
     Rule.Thm("sqrt_square_equation_left_4",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_4}),   
neuper@37950
   349
	      (* c*d/sqrt(x)=b -> c^2*d^2/x=b^2 *)
walther@59871
   350
     Rule.Thm("sqrt_square_equation_left_5",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_5}),   
neuper@37950
   351
	      (* c/d*sqrt(x)=b -> c^2/d^2x=b^2 *)
walther@59871
   352
     Rule.Thm("sqrt_square_equation_left_6",ThmC.numerals_to_Free @{thm sqrt_square_equation_left_6})  
neuper@37950
   353
	      (* c*d/g*sqrt(x)=b -> c^2*d^2/g^2x=b^2 *)
neuper@37950
   354
    ],
walther@59878
   355
    scr = Rule.Empty_Prog
wneuper@59406
   356
   });
neuper@37950
   357
wneuper@59526
   358
\<close> ML \<open>
neuper@37950
   359
(* -- right 28.8.02--*)
neuper@37950
   360
(*isolate the bound variable in an sqrt right equation; 'bdv' is a meta-constant*)
s1210629013@55444
   361
 val r_sqrt_isolate = prep_rls'(
walther@59851
   362
     Rule_Def.Repeat {id = "r_sqrt_isolate", preconds = [], 
neuper@37950
   363
	  rew_ord = ("termlessI",termlessI), 
walther@59851
   364
          erls = RootEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
neuper@37950
   365
     rules = [
walther@59871
   366
     Rule.Thm("sqrt_square_1",ThmC.numerals_to_Free @{thm sqrt_square_1}),
walther@60260
   367
                           (* (sqrt a) \<up> 2 -> a *)
walther@59871
   368
     Rule.Thm("sqrt_square_2",ThmC.numerals_to_Free @{thm sqrt_square_2}), 
walther@60260
   369
                           (* sqrt (a \<up> 2) -> a *)
walther@59871
   370
     Rule.Thm("sqrt_times_root_1",ThmC.numerals_to_Free @{thm sqrt_times_root_1}),
neuper@37950
   371
           (* sqrt a sqrt b -> sqrt(ab) *)
walther@59871
   372
     Rule.Thm("sqrt_times_root_2",ThmC.numerals_to_Free @{thm sqrt_times_root_2}),
neuper@37950
   373
       (* a sqrt b sqrt c -> a sqrt(bc) *)
walther@59871
   374
     Rule.Thm("sqrt_isolate_r_add1",ThmC.numerals_to_Free @{thm sqrt_isolate_r_add1}),
neuper@37950
   375
       (* a= d+e*sqrt(x) -> a-d=e*sqrt(x) *)
walther@59871
   376
     Rule.Thm("sqrt_isolate_r_add2",ThmC.numerals_to_Free @{thm sqrt_isolate_r_add2}),
neuper@37950
   377
       (* a= d+  sqrt(x) -> a-d=  sqrt(x) *)
walther@59871
   378
     Rule.Thm("sqrt_isolate_r_add3",ThmC.numerals_to_Free @{thm sqrt_isolate_r_add3}),
neuper@37950
   379
       (* a=d+e*g/sqrt(x)->a-d=e*g/sqrt(x)*)
walther@59871
   380
     Rule.Thm("sqrt_isolate_r_add4",ThmC.numerals_to_Free @{thm sqrt_isolate_r_add4}),
neuper@37950
   381
       (* a= d+g/sqrt(x) -> a-d=g/sqrt(x) *)
walther@59871
   382
     Rule.Thm("sqrt_isolate_r_add5",ThmC.numerals_to_Free @{thm sqrt_isolate_r_add5}),
neuper@37950
   383
       (* a=d+e*g/h*sqrt(x)->a-d=e*g/h*sqrt(x)*)
walther@59871
   384
     Rule.Thm("sqrt_isolate_r_add6",ThmC.numerals_to_Free @{thm sqrt_isolate_r_add6}),
neuper@37950
   385
       (* a= d+g/h*sqrt(x) -> a-d=g/h*sqrt(x) *)
walther@59871
   386
   (*Rule.Thm("sqrt_isolate_r_div",ThmC.numerals_to_Free @{thm sqrt_isolate_r_div}),*)
neuper@37950
   387
       (* a=e*sqrt(x) -> a/e = sqrt(x) *)
walther@59871
   388
     Rule.Thm("sqrt_square_equation_right_1",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_1}),
neuper@37950
   389
	      (* a=sqrt(x) ->a^2=x *)
walther@59871
   390
     Rule.Thm("sqrt_square_equation_right_2",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_2}),
neuper@37950
   391
	      (* a=c*sqrt(x) ->a^2=c^2*x *)
walther@59871
   392
     Rule.Thm("sqrt_square_equation_right_3",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_3}),
neuper@37950
   393
	      (* a=c/sqrt(x) ->a^2=c^2/x *)
walther@59871
   394
     Rule.Thm("sqrt_square_equation_right_4",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_4}), 
neuper@37950
   395
	      (* a=c*d/sqrt(x) ->a^2=c^2*d^2/x *)
walther@59871
   396
     Rule.Thm("sqrt_square_equation_right_5",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_5}),
neuper@37950
   397
	      (* a=c/e*sqrt(x) ->a^2=c^2/e^2x *)
walther@59871
   398
     Rule.Thm("sqrt_square_equation_right_6",ThmC.numerals_to_Free @{thm sqrt_square_equation_right_6})
neuper@37950
   399
	      (* a=c*d/g*sqrt(x) ->a^2=c^2*d^2/g^2*x *)
neuper@37950
   400
    ],
walther@59878
   401
    scr = Rule.Empty_Prog
wneuper@59406
   402
   });
neuper@37950
   403
wneuper@59526
   404
\<close> ML \<open>
s1210629013@55444
   405
val rooteq_simplify = prep_rls'(
walther@59851
   406
  Rule_Def.Repeat {id = "rooteq_simplify", 
neuper@37950
   407
       preconds = [], rew_ord = ("termlessI",termlessI), 
walther@59851
   408
       erls = RootEq_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59997
   409
       (*asm_thm = [("sqrt_square_1", "")],*)
walther@59871
   410
       rules = [Rule.Thm  ("real_assoc_1",ThmC.numerals_to_Free @{thm real_assoc_1}),
neuper@37950
   411
                             (* a+(b+c) = a+b+c *)
walther@59871
   412
                Rule.Thm  ("real_assoc_2",ThmC.numerals_to_Free @{thm real_assoc_2}),
neuper@37950
   413
                             (* a*(b*c) = a*b*c *)
walther@59878
   414
                Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_"),
walther@59878
   415
                Rule.Eval ("Groups.minus_class.minus", (**)eval_binop "#sub_"),
walther@59878
   416
                Rule.Eval ("Groups.times_class.times", (**)eval_binop "#mult_"),
walther@59878
   417
                Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),
walther@59878
   418
                Rule.Eval ("NthRoot.sqrt", eval_sqrt "#sqrt_"),
walther@60275
   419
                Rule.Eval ("Transcendental.powr" , (**)eval_binop "#power_"),
walther@59871
   420
                Rule.Thm("real_plus_binom_pow2",ThmC.numerals_to_Free @{thm real_plus_binom_pow2}),
walther@59871
   421
                Rule.Thm("real_minus_binom_pow2",ThmC.numerals_to_Free @{thm real_minus_binom_pow2}),
walther@59871
   422
                Rule.Thm("realpow_mul",ThmC.numerals_to_Free @{thm realpow_mul}),    
neuper@37950
   423
                     (* (a * b)^n = a^n * b^n*)
walther@59871
   424
                Rule.Thm("sqrt_times_root_1",ThmC.numerals_to_Free @{thm sqrt_times_root_1}), 
neuper@37950
   425
                     (* sqrt b * sqrt c = sqrt(b*c) *)
walther@59871
   426
                Rule.Thm("sqrt_times_root_2",ThmC.numerals_to_Free @{thm sqrt_times_root_2}),
neuper@37950
   427
                     (* a * sqrt a * sqrt b = a * sqrt(a*b) *)
walther@59871
   428
                Rule.Thm("sqrt_square_2",ThmC.numerals_to_Free @{thm sqrt_square_2}),
walther@60260
   429
                            (* sqrt (a \<up> 2) = a *)
walther@59871
   430
                Rule.Thm("sqrt_square_1",ThmC.numerals_to_Free @{thm sqrt_square_1}) 
walther@60260
   431
                            (* sqrt a  \<up>  2 = a *)
neuper@37950
   432
                ],
walther@59878
   433
       scr = Rule.Empty_Prog
wneuper@59406
   434
    });
wneuper@59472
   435
\<close>
wenzelm@60286
   436
setup_rule
wenzelm@60286
   437
  RootEq_erls = RootEq_erls and
wenzelm@60286
   438
  rooteq_srls = rooteq_srls and
wenzelm@60286
   439
  sqrt_isolate = sqrt_isolate and
wenzelm@60286
   440
  l_sqrt_isolate = l_sqrt_isolate and
wenzelm@60286
   441
  r_sqrt_isolate = r_sqrt_isolate and
wenzelm@60286
   442
  rooteq_simplify = rooteq_simplify
wneuper@59526
   443
wneuper@59526
   444
subsection \<open>problems\<close>
walther@60269
   445
(*Ambiguous input produces 5 parse trees -----------------------------------------------------\\*)
wneuper@59472
   446
setup \<open>KEStore_Elems.add_pbts
s1210629013@55363
   447
  (* ---------root----------- *)
walther@59973
   448
  [(Problem.prep_input thy "pbl_equ_univ_root" [] Problem.id_empty
walther@59997
   449
      (["rootX", "univariate", "equation"],
walther@59997
   450
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55339
   451
          ("#Where" ,["(lhs e_e) is_rootTerm_in  (v_v::real) | " ^
s1210629013@55339
   452
	          "(rhs e_e) is_rootTerm_in  (v_v::real)"]),
s1210629013@55339
   453
          ("#Find"  ,["solutions v_v'i'"])],
s1210629013@55339
   454
        RootEq_prls, SOME "solve (e_e::bool, v_v)", [])),
s1210629013@55339
   455
    (* ---------sqrt----------- *)
walther@59973
   456
    (Problem.prep_input thy "pbl_equ_univ_root_sq" [] Problem.id_empty
walther@59997
   457
      (["sq", "rootX", "univariate", "equation"],
walther@59997
   458
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55339
   459
          ("#Where" ,["( ((lhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
s1210629013@55339
   460
            "  ((lhs e_e) is_normSqrtTerm_in (v_v::real))   )  |" ^
s1210629013@55339
   461
	          "( ((rhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
s1210629013@55339
   462
            "  ((rhs e_e) is_normSqrtTerm_in (v_v::real))   )"]),
s1210629013@55339
   463
          ("#Find"  ,["solutions v_v'i'"])],
walther@59803
   464
          RootEq_prls,  SOME "solve (e_e::bool, v_v)", [["RootEq", "solve_sq_root_equation"]])),
wneuper@59370
   465
    (* ---------normalise----------- *)
walther@59973
   466
    (Problem.prep_input thy "pbl_equ_univ_root_norm" [] Problem.id_empty
walther@59997
   467
      (["normalise", "rootX", "univariate", "equation"],
walther@59997
   468
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55339
   469
          ("#Where" ,["( ((lhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
s1210629013@55339
   470
            "  Not((lhs e_e) is_normSqrtTerm_in (v_v::real)))  | " ^
s1210629013@55339
   471
	          "( ((rhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
s1210629013@55339
   472
            "  Not((rhs e_e) is_normSqrtTerm_in (v_v::real)))"]),
s1210629013@55339
   473
          ("#Find"  ,["solutions v_v'i'"])],
walther@59997
   474
        RootEq_prls, SOME "solve (e_e::bool, v_v)", [["RootEq", "norm_sq_root_equation"]]))]\<close>
walther@60269
   475
(*Ambiguous input produces 5 parse trees -----------------------------------------------------//*)
neuper@37950
   476
wneuper@59526
   477
subsection \<open>methods\<close>
wneuper@59472
   478
setup \<open>KEStore_Elems.add_mets
walther@60154
   479
    [MethodC.prep_input thy "met_rooteq" [] MethodC.id_empty
s1210629013@55373
   480
      (["RootEq"], [],
walther@59852
   481
        {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
wneuper@59545
   482
          crls=RootEq_crls, errpats = [], nrls = norm_Poly}, @{thm refl})]
wneuper@59473
   483
\<close>
wneuper@59370
   484
    (*-- normalise 20.10.02 --*)
wneuper@59504
   485
partial_function (tailrec) norm_sqrt_equ :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   486
  where
walther@59635
   487
"norm_sqrt_equ e_e v_v = (
walther@59635
   488
  let
walther@59635
   489
    e_e = (
walther@59637
   490
      (Repeat(Try (Rewrite ''makex1_x''))) #>
walther@59637
   491
      (Try (Repeat (Rewrite_Set ''expand_rootbinoms''))) #>
walther@59637
   492
      (Try (Rewrite_Set ''rooteq_simplify'')) #>
walther@59637
   493
      (Try (Repeat (Rewrite_Set ''make_rooteq''))) #>
walther@59635
   494
      (Try (Rewrite_Set ''rooteq_simplify''))
walther@59635
   495
      ) e_e
walther@59635
   496
  in
walther@59635
   497
    SubProblem (''RootEq'', [''univariate'', ''equation''], [''no_met''])
walther@59635
   498
      [BOOL e_e, REAL v_v])"
wneuper@59473
   499
setup \<open>KEStore_Elems.add_mets
walther@60154
   500
    [MethodC.prep_input thy "met_rooteq_norm" [] MethodC.id_empty
walther@59997
   501
      (["RootEq", "norm_sq_root_equation"],
walther@59997
   502
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55373
   503
          ("#Where" ,["( ((lhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
s1210629013@55373
   504
              "  Not((lhs e_e) is_normSqrtTerm_in (v_v::real)))  | " ^
s1210629013@55373
   505
              "( ((rhs e_e) is_sqrtTerm_in (v_v::real)) &" ^
s1210629013@55373
   506
              "  Not((rhs e_e) is_normSqrtTerm_in (v_v::real)))"]),
s1210629013@55373
   507
          ("#Find"  ,["solutions v_v'i'"])],
walther@59852
   508
        {rew_ord'="termlessI", rls'=RootEq_erls, srls=Rule_Set.empty, prls=RootEq_prls, calc=[],
s1210629013@55373
   509
          crls=RootEq_crls, errpats = [], nrls = norm_Poly},
wneuper@59551
   510
        @{thm norm_sqrt_equ.simps})]
wneuper@59473
   511
\<close>
wneuper@59545
   512
wneuper@59504
   513
partial_function (tailrec) solve_sqrt_equ :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   514
  where
wneuper@59504
   515
"solve_sqrt_equ e_e v_v =
wneuper@59504
   516
  (let
walther@59635
   517
    e_e = (
walther@59637
   518
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''sqrt_isolate'')) #> 
walther@59637
   519
      (Try (Rewrite_Set         ''rooteq_simplify'')) #> 
walther@59637
   520
      (Try (Repeat (Rewrite_Set ''expand_rootbinoms''))) #>
walther@59803
   521
      (Try (Repeat (Rewrite_Set ''make_rooteq''))) #>
walther@59635
   522
      (Try (Rewrite_Set ''rooteq_simplify'')) ) e_e;
wneuper@59504
   523
    L_L =
wneuper@59504
   524
     (if (((lhs e_e) is_sqrtTerm_in v_v) | ((rhs e_e) is_sqrtTerm_in v_v))               
walther@59635
   525
      then
walther@59635
   526
        SubProblem (''RootEq'', [''normalise'', ''rootX'', ''univariate'', ''equation''], [''no_met''])
walther@59635
   527
          [BOOL e_e, REAL v_v]
walther@59635
   528
      else
walther@59635
   529
        SubProblem (''RootEq'',[''univariate'',''equation''], [''no_met''])
walther@59635
   530
          [BOOL e_e, REAL v_v])
wneuper@59504
   531
  in Check_elementwise L_L {(v_v::real). Assumptions})"
wneuper@59473
   532
setup \<open>KEStore_Elems.add_mets
walther@60154
   533
    [MethodC.prep_input thy "met_rooteq_sq" [] MethodC.id_empty
walther@59803
   534
      (["RootEq", "solve_sq_root_equation"],
s1210629013@55373
   535
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55373
   536
          ("#Where" ,["(((lhs e_e) is_sqrtTerm_in (v_v::real))     & " ^
s1210629013@55373
   537
              " ((lhs e_e) is_normSqrtTerm_in (v_v::real))) |" ^
s1210629013@55373
   538
              "(((rhs e_e) is_sqrtTerm_in (v_v::real))     & " ^
s1210629013@55373
   539
              " ((rhs e_e) is_normSqrtTerm_in (v_v::real)))"]),
s1210629013@55373
   540
          ("#Find"  ,["solutions v_v'i'"])],
s1210629013@55373
   541
        {rew_ord'="termlessI", rls'=RootEq_erls, srls = rooteq_srls, prls = RootEq_prls, calc = [],
s1210629013@55373
   542
          crls=RootEq_crls, errpats = [], nrls = norm_Poly},
wneuper@59551
   543
        @{thm solve_sqrt_equ.simps})]
wneuper@59473
   544
\<close>
s1210629013@55373
   545
    (*-- right 28.08.02 --*)
wneuper@59504
   546
partial_function (tailrec) solve_sqrt_equ_right :: "bool \<Rightarrow> real \<Rightarrow> bool list"
walther@59635
   547
  where
walther@59635
   548
"solve_sqrt_equ_right e_e v_v =
walther@59635
   549
  (let
walther@59635
   550
    e_e = (
walther@59637
   551
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''r_sqrt_isolate'')) #>
walther@59637
   552
      (Try (Rewrite_Set ''rooteq_simplify'')) #>
walther@59637
   553
      (Try (Repeat (Rewrite_Set ''expand_rootbinoms''))) #>
walther@59637
   554
      (Try (Repeat (Rewrite_Set ''make_rooteq''))) #>
walther@59635
   555
      (Try (Rewrite_Set ''rooteq_simplify'')) ) e_e
walther@59635
   556
  in
walther@59635
   557
    if (rhs e_e) is_sqrtTerm_in v_v
walther@59635
   558
    then
walther@59635
   559
      SubProblem (''RootEq'',[''normalise'',''rootX'',''univariate'',''equation''], [''no_met''])
walther@59635
   560
        [BOOL e_e, REAL v_v]
walther@59635
   561
    else
walther@59635
   562
      SubProblem (''RootEq'',[''univariate'', ''equation''], [''no_met''])
walther@59635
   563
        [BOOL e_e, REAL v_v])"
wneuper@59473
   564
setup \<open>KEStore_Elems.add_mets
walther@60154
   565
    [MethodC.prep_input thy "met_rooteq_sq_right" [] MethodC.id_empty
walther@59997
   566
      (["RootEq", "solve_right_sq_root_equation"],
walther@59997
   567
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55373
   568
          ("#Where" ,["(rhs e_e) is_sqrtTerm_in v_v"]),
s1210629013@55373
   569
          ("#Find"  ,["solutions v_v'i'"])],
walther@59852
   570
        {rew_ord' = "termlessI", rls' = RootEq_erls, srls = Rule_Set.empty, prls = RootEq_prls, calc = [],
s1210629013@55373
   571
          crls = RootEq_crls, errpats = [], nrls = norm_Poly},
wneuper@59551
   572
        @{thm solve_sqrt_equ_right.simps})]
wneuper@59473
   573
\<close>
s1210629013@55373
   574
    (*-- left 28.08.02 --*)
wneuper@59504
   575
partial_function (tailrec) solve_sqrt_equ_left :: "bool \<Rightarrow> real \<Rightarrow> bool list"
wneuper@59504
   576
  where
wneuper@59504
   577
"solve_sqrt_equ_left e_e v_v =
walther@59635
   578
  (let
walther@59635
   579
    e_e = (
walther@59637
   580
      (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''l_sqrt_isolate'')) #> 
walther@59637
   581
      (Try (Rewrite_Set ''rooteq_simplify'')) #>
walther@59637
   582
      (Try (Repeat (Rewrite_Set ''expand_rootbinoms''))) #>
walther@59637
   583
      (Try (Repeat (Rewrite_Set ''make_rooteq''))) #>
walther@59635
   584
      (Try (Rewrite_Set ''rooteq_simplify'')) ) e_e
walther@59635
   585
  in
walther@59635
   586
    if (lhs e_e) is_sqrtTerm_in v_v
walther@59635
   587
    then
walther@59635
   588
      SubProblem (''RootEq'', [''normalise'',''rootX'',''univariate'',''equation''], [''no_met''])
walther@59635
   589
        [BOOL e_e, REAL v_v]
walther@59635
   590
    else
walther@59635
   591
      SubProblem (''RootEq'',[''univariate'',''equation''], [''no_met''])
walther@59635
   592
        [BOOL e_e, REAL v_v])"
wneuper@59473
   593
setup \<open>KEStore_Elems.add_mets
walther@60154
   594
    [MethodC.prep_input thy "met_rooteq_sq_left" [] MethodC.id_empty
walther@59997
   595
      (["RootEq", "solve_left_sq_root_equation"],
walther@59997
   596
        [("#Given" ,["equality e_e", "solveFor v_v"]),
s1210629013@55373
   597
          ("#Where" ,["(lhs e_e) is_sqrtTerm_in v_v"]),
s1210629013@55373
   598
          ("#Find"  ,["solutions v_v'i'"])],
walther@59852
   599
        {rew_ord'="termlessI", rls'=RootEq_erls, srls=Rule_Set.empty, prls=RootEq_prls, calc=[],
s1210629013@55373
   600
          crls=RootEq_crls, errpats = [], nrls = norm_Poly},
wneuper@59551
   601
        @{thm solve_sqrt_equ_left.simps})]
wneuper@59473
   602
\<close>
wneuper@59526
   603
ML \<open>
wneuper@59526
   604
\<close> ML \<open>
wneuper@59526
   605
\<close> 
neuper@37906
   606
end