src/Tools/isac/Knowledge/EqSystem.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Wed, 29 May 2019 10:36:16 +0200
changeset 59545 4035ec339062
parent 59505 a1f223658994
child 59551 6ea6d9c377a0
permissions -rw-r--r--
[-Test_Isac] funpack: switch from Script to partial_function
     1 (* equational systems, minimal -- for use in Biegelinie
     2    author: Walther Neuper
     3    050826,
     4    (c) due to copyright terms
     5 *)
     6 
     7 theory EqSystem imports Integrate Rational Root begin
     8 
     9 consts
    10 
    11   occur'_exactly'_in :: 
    12    "[real list, real list, 'a] => bool" ("_ from _ occur'_exactly'_in _")
    13 
    14   (*descriptions in the related problems*)
    15   solveForVars       :: "real list => toreall"
    16   solution           :: "bool list => toreall"
    17 
    18   (*the CAS-command, eg. "solveSystem [x+y=1,y=2] [x,y]"*)
    19   solveSystem        :: "[bool list, real list] => bool list"
    20 
    21   (*Script-names*)
    22   SolveSystemScript  :: "[bool list, real list,     bool list]  
    23 						 => bool list"
    24                   ("((Script SolveSystemScript (_ _ =))// (_))" 9)
    25 
    26 axiomatization where
    27 (*stated as axioms, todo: prove as theorems
    28   'bdv' is a constant handled on the meta-level 
    29    specifically as a 'bound variable'            *)
    30 
    31   commute_0_equality:  "(0 = a) = (a = 0)" and
    32 
    33   (*WN0510 see simliar rules 'isolate_' 'separate_' (by RL)
    34     [bdv_1,bdv_2,bdv_3,bdv_4] work also for 2 and 3 bdvs, ugly !*)
    35   separate_bdvs_add:   
    36     "[| [] from [bdv_1,bdv_2,bdv_3,bdv_4] occur_exactly_in a |] 
    37 		      			     ==> (a + b = c) = (b = c + -1*a)" and
    38   separate_bdvs0:
    39     "[| some_of [bdv_1,bdv_2,bdv_3,bdv_4] occur_in b; Not (b=!=0)  |] 
    40 		      			     ==> (a = b) = (a + -1*b = 0)" and
    41   separate_bdvs_add1:  
    42     "[| some_of [bdv_1,bdv_2,bdv_3,bdv_4] occur_in c |] 
    43 		      			     ==> (a = b + c) = (a + -1*c = b)" and
    44   separate_bdvs_add2:
    45     "[| Not (some_of [bdv_1,bdv_2,bdv_3,bdv_4] occur_in a) |] 
    46 		      			     ==> (a + b = c) = (b = -1*a + c)" and
    47   separate_bdvs_mult:  
    48     "[| [] from [bdv_1,bdv_2,bdv_3,bdv_4] occur_exactly_in a; Not (a=!=0) |] 
    49 		      			     ==>(a * b = c) = (b = c / a)" 
    50 axiomatization where (*..if replaced by "and" we get an error in 
    51   ---  rewrite in [EqSystem,normalise,2x2] --- step "--- 3---";*)
    52   order_system_NxN:     "[a,b] = [b,a]"
    53   (*requires rew_ord for termination, eg. ord_simplify_Integral;
    54     works for lists of any length, interestingly !?!*)
    55 
    56 ML \<open>
    57 val thy = @{theory};
    58 
    59 (** eval functions **)
    60 
    61 (*certain variables of a given list occur _all_ in a term
    62   args: all: ..variables, which are under consideration (eg. the bound vars)
    63         vs:  variables which must be in t, 
    64              and none of the others in all must be in t
    65         t: the term under consideration
    66  *)
    67 fun occur_exactly_in vs all t =
    68     let fun occurs_in' a b = occurs_in b a
    69     in foldl and_ (true, map (occurs_in' t) vs)
    70        andalso not (foldl or_ (false, map (occurs_in' t) 
    71                                           (subtract op = vs all)))
    72     end;
    73 
    74 (*("occur_exactly_in", ("EqSystem.occur'_exactly'_in", 
    75 			eval_occur_exactly_in "#eval_occur_exactly_in_"))*)
    76 fun eval_occur_exactly_in _ "EqSystem.occur'_exactly'_in"
    77 			  (p as (Const ("EqSystem.occur'_exactly'_in",_) 
    78 				       $ vs $ all $ t)) _ =
    79     if occur_exactly_in (TermC.isalist2list vs) (TermC.isalist2list all) t
    80     then SOME ((Rule.term2str p) ^ " = True",
    81 	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
    82     else SOME ((Rule.term2str p) ^ " = False",
    83 	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
    84   | eval_occur_exactly_in _ _ _ _ = NONE;
    85 \<close>
    86 setup \<open>KEStore_Elems.add_calcs
    87   [("occur_exactly_in",
    88 	    ("EqSystem.occur'_exactly'_in",
    89 	      eval_occur_exactly_in "#eval_occur_exactly_in_"))]\<close>
    90 ML \<open>
    91 (** rewrite order 'ord_simplify_System' **)
    92 
    93 (* order wrt. several linear (i.e. without exponents) variables "c","c_2",..
    94    which leaves the monomials containing c, c_2,... at the end of an Integral
    95    and puts the c, c_2,... rightmost within a monomial.
    96 
    97    WN050906 this is a quick and dirty adaption of ord_make_polynomial_in,
    98    which was most adequate, because it uses size_of_term*)
    99 (**)
   100 local (*. for simplify_System .*)
   101 (**)
   102 open Term;  (* for type order = EQUAL | LESS | GREATER *)
   103 
   104 fun pr_ord EQUAL = "EQUAL"
   105   | pr_ord LESS  = "LESS"
   106   | pr_ord GREATER = "GREATER";
   107 
   108 fun dest_hd' (Const (a, T)) = (((a, 0), T), 0)
   109   | dest_hd' (Free (ccc, T)) =
   110     (case Symbol.explode ccc of
   111 	"c"::[] => ((("|||||||||||||||||||||", 0), T), 1)(*greatest string WN*)
   112       | "c"::"_"::_ => ((("|||||||||||||||||||||", 0), T), 1)
   113       | _ => (((ccc, 0), T), 1))
   114   | dest_hd' (Var v) = (v, 2)
   115   | dest_hd' (Bound i) = ((("", i), dummyT), 3)
   116   | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4);
   117 
   118 fun size_of_term' (Free (ccc, _)) =
   119     (case Symbol.explode ccc of (*WN0510 hack for the bound variables*)
   120 	"c"::[] => 1000
   121       | "c"::"_"::is => 1000 * ((TermC.int_of_str o implode) is)
   122       | _ => 1)
   123   | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
   124   | size_of_term' (f$t) = size_of_term' f  +  size_of_term' t
   125   | size_of_term' _ = 1;
   126 
   127 fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
   128     (case term_ord' pr thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord)
   129   | term_ord' pr thy (t, u) =
   130     (if pr
   131      then 
   132        let
   133          val (f, ts) = strip_comb t and (g, us) = strip_comb u;
   134          val _ = tracing ("t= f@ts= \"" ^ Rule.term_to_string''' thy f ^ "\" @ \"[" ^
   135            commas (map (Rule.term_to_string''' thy) ts) ^ "]\"");
   136          val _ = tracing ("u= g@us= \"" ^ Rule.term_to_string''' thy g ^ "\" @ \"[" ^
   137            commas (map (Rule.term_to_string''' thy) us) ^ "]\"");
   138          val _ = tracing ("size_of_term(t,u)= (" ^ string_of_int (size_of_term' t) ^ ", " ^
   139            string_of_int (size_of_term' u) ^ ")");
   140          val _ = tracing ("hd_ord(f,g)      = " ^ ((pr_ord o hd_ord) (f,g)));
   141          val _ = tracing ("terms_ord (ts,us) = " ^(pr_ord o terms_ord str false) (ts,us));
   142          val _=tracing("-------");
   143        in () end
   144      else ();
   145     case int_ord (size_of_term' t, size_of_term' u) of
   146       EQUAL =>
   147         let val (f, ts) = strip_comb t and (g, us) = strip_comb u 
   148         in (case hd_ord (f, g) of 
   149               EQUAL => (terms_ord str pr) (ts, us) 
   150             | ord => ord)
   151         end
   152 	 | ord => ord)
   153 and hd_ord (f, g) =                                        (* ~ term.ML *)
   154   prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord (dest_hd' f, dest_hd' g)
   155 and terms_ord str pr (ts, us) = list_ord (term_ord' pr (Celem.assoc_thy "Isac"))(ts, us);
   156 (**)
   157 in
   158 (**)
   159 (*WN0510 for preliminary use in eval_order_system, see case-study mat-eng.tex
   160 fun ord_simplify_System_rev (pr:bool) thy subst tu = 
   161     (term_ord' pr thy (Library.swap tu) = LESS);*)
   162 
   163 (*for the rls's*)
   164 fun ord_simplify_System (pr:bool) thy subst tu = 
   165     (term_ord' pr thy tu = LESS);
   166 (**)
   167 end;
   168 (**)
   169 Rule.rew_ord' := overwritel (! Rule.rew_ord',
   170 [("ord_simplify_System", ord_simplify_System false thy)
   171  ]);
   172 \<close>
   173 ML \<open>
   174 (** rulesets **)
   175 
   176 (*.adapted from 'order_add_mult_in' by just replacing the rew_ord.*)
   177 val order_add_mult_System = 
   178   Rule.Rls{id = "order_add_mult_System", preconds = [], 
   179       rew_ord = ("ord_simplify_System",
   180 		 ord_simplify_System false @{theory "Integrate"}),
   181       erls = Rule.e_rls,srls = Rule.Erls, calc = [], errpatts = [],
   182       rules = [Rule.Thm ("mult_commute",TermC.num_str @{thm mult.commute}),
   183 	       (* z * w = w * z *)
   184 	       Rule.Thm ("real_mult_left_commute",TermC.num_str @{thm real_mult_left_commute}),
   185 	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
   186 	       Rule.Thm ("mult_assoc",TermC.num_str @{thm mult.assoc}),		
   187 	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
   188 	       Rule.Thm ("add_commute",TermC.num_str @{thm add.commute}),	
   189 	       (*z + w = w + z*)
   190 	       Rule.Thm ("add_left_commute",TermC.num_str @{thm add.left_commute}),
   191 	       (*x + (y + z) = y + (x + z)*)
   192 	       Rule.Thm ("add_assoc",TermC.num_str @{thm add.assoc})	               
   193 	       (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
   194 	       ], 
   195       scr = Rule.EmptyScr};
   196 \<close>
   197 ML \<open>
   198 (*.adapted from 'norm_Rational' by
   199   #1 using 'ord_simplify_System' in 'order_add_mult_System'
   200   #2 NOT using common_nominator_p                          .*)
   201 val norm_System_noadd_fractions = 
   202   Rule.Rls {id = "norm_System_noadd_fractions", preconds = [], 
   203        rew_ord = ("dummy_ord",Rule.dummy_ord), 
   204        erls = norm_rat_erls, srls = Rule.Erls, calc = [], errpatts = [],
   205        rules = [(*sequence given by operator precedence*)
   206 		Rule.Rls_ discard_minus,
   207 		Rule.Rls_ powers,
   208 		Rule.Rls_ rat_mult_divide,
   209 		Rule.Rls_ expand,
   210 		Rule.Rls_ reduce_0_1_2,
   211 		Rule.Rls_ (*order_add_mult #1*) order_add_mult_System,
   212 		Rule.Rls_ collect_numerals,
   213 		(*Rule.Rls_ add_fractions_p, #2*)
   214 		Rule.Rls_ cancel_p
   215 		],
   216        scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
   217        };
   218 \<close>
   219 ML \<open>
   220 (*.adapted from 'norm_Rational' by
   221   *1* using 'ord_simplify_System' in 'order_add_mult_System'.*)
   222 val norm_System = 
   223   Rule.Rls {id = "norm_System", preconds = [], 
   224        rew_ord = ("dummy_ord",Rule.dummy_ord), 
   225        erls = norm_rat_erls, srls = Rule.Erls, calc = [], errpatts = [],
   226        rules = [(*sequence given by operator precedence*)
   227 		Rule.Rls_ discard_minus,
   228 		Rule.Rls_ powers,
   229 		Rule.Rls_ rat_mult_divide,
   230 		Rule.Rls_ expand,
   231 		Rule.Rls_ reduce_0_1_2,
   232 		Rule.Rls_ (*order_add_mult *1*) order_add_mult_System,
   233 		Rule.Rls_ collect_numerals,
   234 		Rule.Rls_ add_fractions_p,
   235 		Rule.Rls_ cancel_p
   236 		],
   237        scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
   238        };
   239 \<close>
   240 ML \<open>
   241 (*.simplify an equational system BEFORE solving it such that parentheses are
   242    ( ((u0*v0)*w0) + ( ((u1*v1)*w1) * c + ... +((u4*v4)*w4) * c_4 ) )
   243 ATTENTION: works ONLY for bound variables c, c_1, c_2, c_3, c_4 :ATTENTION
   244    This is a copy from 'make_ratpoly_in' with respective reductions:
   245    *0* expand the term, ie. distribute * and / over +
   246    *1* ord_simplify_System instead of termlessI
   247    *2* no add_fractions_p (= common_nominator_p_rls !)
   248    *3* discard_parentheses only for (.*(.*.))
   249    analoguous to simplify_Integral                                       .*)
   250 val simplify_System_parenthesized = 
   251   Rule.Seq {id = "simplify_System_parenthesized", preconds = []:term list, 
   252        rew_ord = ("dummy_ord", Rule.dummy_ord),
   253       erls = Atools_erls, srls = Rule.Erls, calc = [], errpatts = [],
   254       rules = [Rule.Thm ("distrib_right",TermC.num_str @{thm distrib_right}),
   255  	       (*"(?z1.0 + ?z2.0) * ?w = ?z1.0 * ?w + ?z2.0 * ?w"*)
   256 	       Rule.Thm ("add_divide_distrib",TermC.num_str @{thm add_divide_distrib}),
   257  	       (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)
   258 	       (*^^^^^ *0* ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^*)
   259 	       Rule.Rls_ norm_Rational_noadd_fractions(**2**),
   260 	       Rule.Rls_ (*order_add_mult_in*) norm_System_noadd_fractions (**1**),
   261 	       Rule.Thm ("sym_mult_assoc",
   262                      TermC.num_str (@{thm mult.assoc} RS @{thm sym}))
   263 	       (*Rule.Rls_ discard_parentheses *3**),
   264 	       Rule.Rls_ collect_bdv, (*from make_polynomial_in WN051031 welldone?*)
   265 	       Rule.Rls_ separate_bdv2,
   266 	       Rule.Calc ("Rings.divide_class.divide"  ,eval_cancel "#divide_e")
   267 	       ],
   268       scr = Rule.EmptyScr};      
   269 \<close>
   270 ML \<open>
   271 (*.simplify an equational system AFTER solving it;
   272    This is a copy of 'make_ratpoly_in' with the differences
   273    *1* ord_simplify_System instead of termlessI           .*)
   274 (*TODO.WN051031 ^^^^^^^^^^ should be in EACH rls contained *)
   275 val simplify_System = 
   276   Rule.Seq {id = "simplify_System", preconds = []:term list, 
   277        rew_ord = ("dummy_ord", Rule.dummy_ord),
   278       erls = Atools_erls, srls = Rule.Erls, calc = [], errpatts = [],
   279       rules = [Rule.Rls_ norm_Rational,
   280 	       Rule.Rls_ (*order_add_mult_in*) norm_System (**1**),
   281 	       Rule.Rls_ discard_parentheses,
   282 	       Rule.Rls_ collect_bdv, (*from make_polynomial_in WN051031 welldone?*)
   283 	       Rule.Rls_ separate_bdv2,
   284 	       Rule.Calc ("Rings.divide_class.divide"  ,eval_cancel "#divide_e")
   285 	       ],
   286       scr = Rule.EmptyScr};      
   287 (*
   288 val simplify_System = 
   289     Rule.append_rls "simplify_System" simplify_System_parenthesized
   290 	       [Rule.Thm ("sym_add_assoc",
   291                       TermC.num_str (@{thm add.assoc} RS @{thm sym}))];
   292 *)
   293 \<close>
   294 ML \<open>
   295 val isolate_bdvs = 
   296     Rule.Rls {id="isolate_bdvs", preconds = [], 
   297 	 rew_ord = ("e_rew_ord", Rule.e_rew_ord), 
   298 	 erls = Rule.append_rls "erls_isolate_bdvs" Rule.e_rls 
   299 			   [(Rule.Calc ("EqSystem.occur'_exactly'_in", 
   300 				   eval_occur_exactly_in 
   301 				       "#eval_occur_exactly_in_"))
   302 			    ], 
   303 			   srls = Rule.Erls, calc = [], errpatts = [],
   304 	      rules = 
   305              [Rule.Thm ("commute_0_equality", TermC.num_str @{thm commute_0_equality}),
   306 	      Rule.Thm ("separate_bdvs_add", TermC.num_str @{thm separate_bdvs_add}),
   307 	      Rule.Thm ("separate_bdvs_mult", TermC.num_str @{thm separate_bdvs_mult})],
   308 	      scr = Rule.EmptyScr};
   309 \<close>
   310 ML \<open>
   311 val isolate_bdvs_4x4 = 
   312     Rule.Rls {id="isolate_bdvs_4x4", preconds = [], 
   313 	 rew_ord = ("e_rew_ord", Rule.e_rew_ord), 
   314 	 erls = Rule.append_rls 
   315 		    "erls_isolate_bdvs_4x4" Rule.e_rls 
   316 		    [Rule.Calc ("EqSystem.occur'_exactly'_in", 
   317 			   eval_occur_exactly_in "#eval_occur_exactly_in_"),
   318 		     Rule.Calc ("Atools.ident",eval_ident "#ident_"),
   319 		     Rule.Calc ("Atools.some'_occur'_in", 
   320 			   eval_some_occur_in "#some_occur_in_"),
   321                      Rule.Thm ("not_true",TermC.num_str @{thm not_true}),
   322 		     Rule.Thm ("not_false",TermC.num_str @{thm not_false})
   323 			    ], 
   324 	 srls = Rule.Erls, calc = [], errpatts = [],
   325 	 rules = [Rule.Thm ("commute_0_equality", TermC.num_str @{thm commute_0_equality}),
   326 		  Rule.Thm ("separate_bdvs0", TermC.num_str @{thm separate_bdvs0}),
   327 		  Rule.Thm ("separate_bdvs_add1", TermC.num_str @{thm separate_bdvs_add1}),
   328 		  Rule.Thm ("separate_bdvs_add1", TermC.num_str @{thm separate_bdvs_add2}),
   329 		  Rule.Thm ("separate_bdvs_mult", TermC.num_str @{thm separate_bdvs_mult})
   330                  ], scr = Rule.EmptyScr};
   331 
   332 \<close>
   333 ML \<open>
   334 
   335 (*.order the equations in a system such, that a triangular system (if any)
   336    appears as [..c_4 = .., ..., ..., ..c_1 + ..c_2 + ..c_3 ..c_4 = ..].*)
   337 val order_system = 
   338     Rule.Rls {id="order_system", preconds = [], 
   339 	 rew_ord = ("ord_simplify_System", 
   340 		    ord_simplify_System false thy), 
   341 	 erls = Rule.Erls, srls = Rule.Erls, calc = [], errpatts = [],
   342 	 rules = [Rule.Thm ("order_system_NxN", TermC.num_str @{thm order_system_NxN})
   343 		  ],
   344 	 scr = Rule.EmptyScr};
   345 
   346 val prls_triangular = 
   347     Rule.Rls {id="prls_triangular", preconds = [], 
   348 	 rew_ord = ("e_rew_ord", Rule.e_rew_ord), 
   349 	 erls = Rule.Rls {id="erls_prls_triangular", preconds = [], 
   350 		     rew_ord = ("e_rew_ord", Rule.e_rew_ord), 
   351 		     erls = Rule.Erls, srls = Rule.Erls, calc = [], errpatts = [],
   352 		     rules = [(*for precond NTH_CONS ...*)
   353 			      Rule.Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   354 			      Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_")
   355 			      (*immediately repeated rewrite pushes
   356 					    '+' into precondition !*)
   357 			      ],
   358 		     scr = Rule.EmptyScr}, 
   359 	 srls = Rule.Erls, calc = [], errpatts = [],
   360 	 rules = [Rule.Thm ("NTH_CONS",TermC.num_str @{thm NTH_CONS}),
   361 		  Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   362 		  Rule.Thm ("NTH_NIL",TermC.num_str @{thm NTH_NIL}),
   363 		  Rule.Thm ("tl_Cons",TermC.num_str @{thm tl_Cons}),
   364 		  Rule.Thm ("tl_Nil",TermC.num_str @{thm tl_Nil}),
   365 		  Rule.Calc ("EqSystem.occur'_exactly'_in", 
   366 			eval_occur_exactly_in 
   367 			    "#eval_occur_exactly_in_")
   368 		  ],
   369 	 scr = Rule.EmptyScr};
   370 \<close>
   371 ML \<open>
   372 
   373 (*WN060914 quickly created for 4x4; 
   374  more similarity to prls_triangular desirable*)
   375 val prls_triangular4 = 
   376     Rule.Rls {id="prls_triangular4", preconds = [], 
   377 	 rew_ord = ("e_rew_ord", Rule.e_rew_ord), 
   378 	 erls = Rule.Rls {id="erls_prls_triangular4", preconds = [], 
   379 		     rew_ord = ("e_rew_ord", Rule.e_rew_ord), 
   380 		     erls = Rule.Erls, srls = Rule.Erls, calc = [], errpatts = [],
   381 		     rules = [(*for precond NTH_CONS ...*)
   382 			      Rule.Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   383 			      Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_")
   384 			      (*immediately repeated rewrite pushes
   385 					    '+' into precondition !*)
   386 			      ],
   387 		     scr = Rule.EmptyScr}, 
   388 	 srls = Rule.Erls, calc = [], errpatts = [],
   389 	 rules = [Rule.Thm ("NTH_CONS",TermC.num_str @{thm NTH_CONS}),
   390 		  Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   391 		  Rule.Thm ("NTH_NIL",TermC.num_str @{thm NTH_NIL}),
   392 		  Rule.Thm ("tl_Cons",TermC.num_str @{thm tl_Cons}),
   393 		  Rule.Thm ("tl_Nil",TermC.num_str @{thm tl_Nil}),
   394 		  Rule.Calc ("EqSystem.occur'_exactly'_in", 
   395 			eval_occur_exactly_in 
   396 			    "#eval_occur_exactly_in_")
   397 		  ],
   398 	 scr = Rule.EmptyScr};
   399 \<close>
   400 
   401 setup \<open>KEStore_Elems.add_rlss 
   402   [("simplify_System_parenthesized",
   403     (Context.theory_name @{theory}, prep_rls' simplify_System_parenthesized)), 
   404   ("simplify_System", (Context.theory_name @{theory}, prep_rls' simplify_System)), 
   405   ("isolate_bdvs", (Context.theory_name @{theory}, prep_rls' isolate_bdvs)), 
   406   ("isolate_bdvs_4x4", (Context.theory_name @{theory}, prep_rls' isolate_bdvs_4x4)), 
   407   ("order_system", (Context.theory_name @{theory}, prep_rls' order_system)), 
   408   ("order_add_mult_System", (Context.theory_name @{theory}, prep_rls' order_add_mult_System)), 
   409   ("norm_System_noadd_fractions", 
   410     (Context.theory_name @{theory}, prep_rls' norm_System_noadd_fractions)), 
   411   ("norm_System", (Context.theory_name @{theory}, prep_rls' norm_System))]\<close>
   412 
   413 (** problems **)
   414 setup \<open>KEStore_Elems.add_pbts
   415   [(Specify.prep_pbt thy "pbl_equsys" [] Celem.e_pblID
   416       (["system"],
   417         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   418           ("#Find"  ,["solution ss'''"](*''' is copy-named*))],
   419         Rule.append_rls "e_rls" Rule.e_rls [(*for preds in where_*)], SOME "solveSystem e_s v_s", [])),
   420     (Specify.prep_pbt thy "pbl_equsys_lin" [] Celem.e_pblID
   421       (["LINEAR", "system"],
   422         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   423           (*TODO.WN050929 check linearity*)
   424           ("#Find"  ,["solution ss'''"])],
   425         Rule.append_rls "e_rls" Rule.e_rls [(*for preds in where_*)], SOME "solveSystem e_s v_s", [])),
   426     (Specify.prep_pbt thy "pbl_equsys_lin_2x2" [] Celem.e_pblID
   427       (["2x2", "LINEAR", "system"],
   428       (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   429         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   430           ("#Where"  ,["LENGTH (e_s:: bool list) = 2", "LENGTH v_s = 2"]),
   431           ("#Find"  ,["solution ss'''"])],
   432         Rule.append_rls "prls_2x2_linear_system" Rule.e_rls 
   433 			    [Rule.Thm ("LENGTH_CONS",TermC.num_str @{thm LENGTH_CONS}),
   434 			      Rule.Thm ("LENGTH_NIL",TermC.num_str @{thm LENGTH_NIL}),
   435 			      Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   436 			      Rule.Calc ("HOL.eq",eval_equal "#equal_")], 
   437         SOME "solveSystem e_s v_s", [])),
   438     (Specify.prep_pbt thy "pbl_equsys_lin_2x2_tri" [] Celem.e_pblID
   439       (["triangular", "2x2", "LINEAR", "system"],
   440         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   441           ("#Where",
   442             ["(tl v_s) from v_s occur_exactly_in (NTH 1 (e_s::bool list))",
   443               "    v_s  from v_s occur_exactly_in (NTH 2 (e_s::bool list))"]),
   444           ("#Find"  ,["solution ss'''"])],
   445         prls_triangular, SOME "solveSystem e_s v_s", [["EqSystem","top_down_substitution","2x2"]])),
   446     (Specify.prep_pbt thy "pbl_equsys_lin_2x2_norm" [] Celem.e_pblID
   447       (["normalise", "2x2", "LINEAR", "system"],
   448         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   449           ("#Find"  ,["solution ss'''"])],
   450       Rule.append_rls "e_rls" Rule.e_rls [(*for preds in where_*)], 
   451       SOME "solveSystem e_s v_s", 
   452       [["EqSystem","normalise","2x2"]])),
   453     (Specify.prep_pbt thy "pbl_equsys_lin_3x3" [] Celem.e_pblID
   454       (["3x3", "LINEAR", "system"],
   455         (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   456         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   457           ("#Where"  ,["LENGTH (e_s:: bool list) = 3", "LENGTH v_s = 3"]),
   458           ("#Find"  ,["solution ss'''"])],
   459         Rule.append_rls "prls_3x3_linear_system" Rule.e_rls 
   460 			    [Rule.Thm ("LENGTH_CONS",TermC.num_str @{thm LENGTH_CONS}),
   461 			      Rule.Thm ("LENGTH_NIL",TermC.num_str @{thm LENGTH_NIL}),
   462 			      Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   463 			      Rule.Calc ("HOL.eq",eval_equal "#equal_")],
   464         SOME "solveSystem e_s v_s", [])),
   465     (Specify.prep_pbt thy "pbl_equsys_lin_4x4" [] Celem.e_pblID
   466       (["4x4", "LINEAR", "system"],
   467         (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   468         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   469           ("#Where"  ,["LENGTH (e_s:: bool list) = 4", "LENGTH v_s = 4"]),
   470           ("#Find"  ,["solution ss'''"])],
   471         Rule.append_rls "prls_4x4_linear_system" Rule.e_rls 
   472 			    [Rule.Thm ("LENGTH_CONS",TermC.num_str @{thm LENGTH_CONS}),
   473 			      Rule.Thm ("LENGTH_NIL",TermC.num_str @{thm LENGTH_NIL}),
   474 			      Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   475 			      Rule.Calc ("HOL.eq",eval_equal "#equal_")],
   476         SOME "solveSystem e_s v_s", [])),
   477     (Specify.prep_pbt thy "pbl_equsys_lin_4x4_tri" [] Celem.e_pblID
   478       (["triangular", "4x4", "LINEAR", "system"],
   479         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   480           ("#Where" , (*accepts missing variables up to diagional form*)
   481             ["(NTH 1 (v_s::real list)) occurs_in (NTH 1 (e_s::bool list))",
   482               "(NTH 2 (v_s::real list)) occurs_in (NTH 2 (e_s::bool list))",
   483               "(NTH 3 (v_s::real list)) occurs_in (NTH 3 (e_s::bool list))",
   484               "(NTH 4 (v_s::real list)) occurs_in (NTH 4 (e_s::bool list))"]),
   485           ("#Find"  ,["solution ss'''"])],
   486       Rule.append_rls "prls_tri_4x4_lin_sys" prls_triangular
   487 	      [Rule.Calc ("Atools.occurs'_in",eval_occurs_in "")], 
   488       SOME "solveSystem e_s v_s", 
   489       [["EqSystem","top_down_substitution","4x4"]])),
   490     (Specify.prep_pbt thy "pbl_equsys_lin_4x4_norm" [] Celem.e_pblID
   491       (["normalise", "4x4", "LINEAR", "system"],
   492         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   493           (*LENGTH is checked 1 level above*)
   494           ("#Find"  ,["solution ss'''"])],
   495         Rule.append_rls "e_rls" Rule.e_rls [(*for preds in where_*)], 
   496         SOME "solveSystem e_s v_s", 
   497         [["EqSystem","normalise","4x4"]]))]\<close>
   498 
   499 ML \<open>
   500 (*this is for NTH only*)
   501 val srls = Rule.Rls {id="srls_normalise_4x4", 
   502 		preconds = [], 
   503 		rew_ord = ("termlessI",termlessI), 
   504 		erls = Rule.append_rls "erls_in_srls_IntegrierenUnd.." Rule.e_rls
   505 				  [(*for asm in NTH_CONS ...*)
   506 				   Rule.Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   507 				   (*2nd NTH_CONS pushes n+-1 into asms*)
   508 				   Rule.Calc("Groups.plus_class.plus", eval_binop "#add_")
   509 				   ], 
   510 		srls = Rule.Erls, calc = [], errpatts = [],
   511 		rules = [Rule.Thm ("NTH_CONS",TermC.num_str @{thm NTH_CONS}),
   512 			 Rule.Calc("Groups.plus_class.plus", eval_binop "#add_"),
   513 			 Rule.Thm ("NTH_NIL",TermC.num_str @{thm NTH_NIL})],
   514 		scr = Rule.EmptyScr};
   515 \<close>
   516 
   517 (**methods**)
   518 setup \<open>KEStore_Elems.add_mets
   519     [Specify.prep_met thy "met_eqsys" [] Celem.e_metID
   520 	    (["EqSystem"], [],
   521 	      {rew_ord'="tless_true", rls' = Rule.Erls, calc = [], srls = Rule.Erls, prls = Rule.Erls, crls = Rule.Erls,
   522           errpats = [], nrls = Rule.Erls},
   523 	      @{thm refl}),
   524     Specify.prep_met thy "met_eqsys_topdown" [] Celem.e_metID
   525       (["EqSystem","top_down_substitution"], [],
   526         {rew_ord'="tless_true", rls' = Rule.Erls, calc = [], srls = Rule.Erls, prls = Rule.Erls, crls = Rule.Erls,
   527           errpats = [], nrls = Rule.Erls},
   528        @{thm refl})]
   529 \<close>
   530 
   531 partial_function (tailrec) solve_system :: "bool list => real list => bool list"
   532   where
   533 "solve_system e_s v_s =                          
   534   (let e_1 = Take (hd e_s);                                                         
   535        e_1 = ((Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))] 
   536                                   ''isolate_bdvs'' False)) @@                   
   537                (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]
   538                                   ''simplify_System'' False))) e_1;                 
   539        e_2 = Take (hd (tl e_s));                                                    
   540        e_2 = ((Substitute [e_1]) @@                                                 
   541                (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]
   542                                   ''simplify_System_parenthesized'' False)) @@      
   543                (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]
   544                                   ''isolate_bdvs'' False)) @@                   
   545                (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]
   546                                   ''simplify_System'' False))) e_2;                 
   547        e__s = Take [e_1, e_2]                                                       
   548    in Try (Rewrite_Set ''order_system'' False) e__s)                              "
   549 setup \<open>KEStore_Elems.add_mets
   550     [Specify.prep_met thy "met_eqsys_topdown_2x2" [] Celem.e_metID
   551       (["EqSystem", "top_down_substitution", "2x2"],
   552         [("#Given", ["equalities e_s", "solveForVars v_s"]),
   553           ("#Where",
   554             ["(tl v_s) from v_s occur_exactly_in (NTH 1 (e_s::bool list))",
   555               "    v_s  from v_s occur_exactly_in (NTH 2 (e_s::bool list))"]),
   556           ("#Find"  ,["solution ss'''"])],
   557 	      {rew_ord'="ord_simplify_System", rls' = Rule.Erls, calc = [], 
   558 	        srls = Rule.append_rls "srls_top_down_2x2" Rule.e_rls
   559 				      [Rule.Thm ("hd_thm",TermC.num_str @{thm hd_thm}),
   560 				        Rule.Thm ("tl_Cons",TermC.num_str @{thm tl_Cons}),
   561 				        Rule.Thm ("tl_Nil",TermC.num_str @{thm tl_Nil})], 
   562 	        prls = prls_triangular, crls = Rule.Erls, errpats = [], nrls = Rule.Erls},
   563 	      @{thm solve_system.simps}
   564 	    (*"Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   565           "  (let e_1 = Take (hd e_s);                                                " ^
   566           "       e_1 = ((Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   567           "                                  ''isolate_bdvs'' False))     @@               " ^
   568           "               (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   569           "                                  ''simplify_System'' False))) e_1;            " ^
   570           "       e_2 = Take (hd (tl e_s));                                           " ^
   571           "       e_2 = ((Substitute [e_1]) @@                                       " ^
   572           "               (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   573           "                                  ''simplify_System_parenthesized'' False)) @@  " ^
   574           "               (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   575           "                                  ''isolate_bdvs'' False))     @@               " ^
   576           "               (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   577           "                                  ''simplify_System'' False))) e_2;            " ^
   578           "       e__s = Take [e_1, e_2]                                             " ^
   579           "   in (Try (Rewrite_Set ''order_system'' False)) e__s)"*)
   580           (*---------------------------------------------------------------------------
   581             this script does NOT separate the equations as above, 
   582             but it does not yet work due to preliminary script-interpreter,
   583             see eqsystem.sml 'script [EqSystem,top_down_substitution,2x2] Vers.2'
   584           
   585           "Script SolveSystemScript (e_s::bool list) (v_s::real list) =         " ^
   586           "  (let e__s = Take e_s;                                              " ^
   587           "       e_1 = hd e__s;                                               " ^
   588           "       e_2 = hd (tl e__s);                                          " ^
   589           "       e__s = [e_1, Substitute [e_1] e_2]                         " ^
   590           "   in ((Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   591           "                                  ''simplify_System_parenthesized'' False)) @@ " ^
   592           "       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))] " ^
   593           "                              ''isolate_bdvs'' False))              @@   " ^
   594           "       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   595           "                                  ''simplify_System'' False))) e__s)"
   596           ---------------------------------------------------------------------------*))]
   597 \<close>
   598 setup \<open>KEStore_Elems.add_mets
   599     [Specify.prep_met thy "met_eqsys_norm" [] Celem.e_metID
   600 	    (["EqSystem", "normalise"], [],
   601 	      {rew_ord'="tless_true", rls' = Rule.Erls, calc = [], srls = Rule.Erls, prls = Rule.Erls, crls = Rule.Erls,
   602           errpats = [], nrls = Rule.Erls},
   603 	      @{thm refl})]
   604 \<close>
   605 
   606 partial_function (tailrec) solve_system2 :: "bool list => real list => bool list"
   607   where
   608 "solve_system2 e_s v_s =
   609   (let e__s = ((Try (Rewrite_Set ''norm_Rational'' False)) @@
   610                (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))]
   611                                   ''simplify_System_parenthesized'' False)) @@
   612                (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))]
   613                                   ''isolate_bdvs'' False)) @@
   614                (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))]
   615                                   ''simplify_System_parenthesized'' False)) @@
   616                (Try (Rewrite_Set ''order_system'' False))) e_s
   617    in (SubProblem (''EqSystem'', [''LINEAR'', ''system''], [''no_met''])
   618                   [BOOL_LIST e__s, REAL_LIST v_s]))"
   619 setup \<open>KEStore_Elems.add_mets
   620     [Specify.prep_met thy "met_eqsys_norm_2x2" [] Celem.e_metID
   621 	    (["EqSystem","normalise","2x2"],
   622 	      [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   623 		      ("#Find"  ,["solution ss'''"])],
   624 	      {rew_ord'="tless_true", rls' = Rule.Erls, calc = [], 
   625 	        srls = Rule.append_rls "srls_normalise_2x2" Rule.e_rls
   626 				      [Rule.Thm ("hd_thm",TermC.num_str @{thm hd_thm}),
   627 				        Rule.Thm ("tl_Cons",TermC.num_str @{thm tl_Cons}),
   628 				        Rule.Thm ("tl_Nil",TermC.num_str @{thm tl_Nil})], 
   629 		      prls = Rule.Erls, crls = Rule.Erls, errpats = [], nrls = Rule.Erls},
   630 		    @{thm solve_system2.simps}
   631 	    (*"Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   632           "  (let e__s = ((Try (Rewrite_Set ''norm_Rational'' False)) @@                   " ^
   633           "               (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   634           "                                  ''simplify_System_parenthesized'' False)) @@  " ^
   635           "               (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   636           "                                                    ''isolate_bdvs'' False)) @@ " ^
   637           "               (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s),(''bdv_2'', hd (tl v_s))]" ^
   638           "                                  ''simplify_System_parenthesized'' False)) @@  " ^
   639           "               (Try (Rewrite_Set ''order_system'' False))) e_s                  " ^
   640           "   in (SubProblem (''EqSystem'',[''LINEAR'',''system''],[''no_met''])                      " ^
   641           "                  [BOOL_LIST e__s, REAL_LIST v_s]))"*))]
   642 \<close>
   643 
   644 partial_function (tailrec) solve_system3 :: "bool list => real list => bool list"
   645   where
   646 "solve_system3 e_s v_s =
   647   (let e__s =
   648      ((Try (Rewrite_Set ''norm_Rational'' False)) @@
   649       (Repeat (Rewrite ''commute_0_equality'' False)) @@
   650       (Try (Rewrite_Set_Inst [(''bdv_1'', NTH 1 v_s),(''bdv_2'', NTH 2 v_s ),
   651                               (''bdv_3'', NTH 3 v_s),(''bdv_3'', NTH 4 v_s )]
   652               ''simplify_System_parenthesized'' False)) @@
   653       (Try (Rewrite_Set_Inst [(''bdv_1'', NTH 1 v_s),(''bdv_2'', NTH 2 v_s ),
   654                               (''bdv_3'', NTH 3 v_s),(''bdv_3'', NTH 4 v_s )]
   655               ''isolate_bdvs_4x4'' False)) @@
   656       (Try (Rewrite_Set_Inst [(''bdv_1'', NTH 1 v_s),(''bdv_2'', NTH 2 v_s ),
   657                               (''bdv_3'', NTH 3 v_s),(''bdv_3'', NTH 4 v_s )]
   658                ''simplify_System_parenthesized'' False)) @@
   659       (Try (Rewrite_Set ''order_system'' False)))  e_s
   660    in (SubProblem (''EqSystem'', [''LINEAR'', ''system''], [''no_met''])
   661                   [BOOL_LIST e__s, REAL_LIST v_s]))"
   662 setup \<open>KEStore_Elems.add_mets
   663     [Specify.prep_met thy "met_eqsys_norm_4x4" [] Celem.e_metID
   664 	      (["EqSystem","normalise","4x4"],
   665 	       [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   666 	         ("#Find"  ,["solution ss'''"])],
   667 	       {rew_ord'="tless_true", rls' = Rule.Erls, calc = [], 
   668 	         srls = Rule.append_rls "srls_normalise_4x4" srls
   669 	             [Rule.Thm ("hd_thm",TermC.num_str @{thm hd_thm}),
   670 	               Rule.Thm ("tl_Cons",TermC.num_str @{thm tl_Cons}),
   671 	               Rule.Thm ("tl_Nil",TermC.num_str @{thm tl_Nil})], 
   672 		       prls = Rule.Erls, crls = Rule.Erls, errpats = [], nrls = Rule.Erls},
   673 		     (*STOPPED.WN06? met ["EqSystem","normalise","4x4"] @@@@@@@@@@@@@@@@@@@@@@@@@@@*)
   674 		     @{thm solve_system3.simps}
   675 	    (*""Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   676            "  (let e__s =                                                               " ^
   677            "     ((Try (Rewrite_Set ''norm_Rational'' False)) @@                            " ^
   678            "      (Repeat (Rewrite ''commute_0_equality'' False)) @@                        " ^
   679            "      (Try (Rewrite_Set_Inst [(''bdv_1'', NTH 1 v_s),(''bdv_2'', NTH 2 v_s ),     " ^
   680            "                              (''bdv_3'', NTH 3 v_s),(''bdv_3'', NTH 4 v_s )]     " ^
   681            "                             ''simplify_System_parenthesized'' False))    @@    " ^
   682            "      (Try (Rewrite_Set_Inst [(''bdv_1'', NTH 1 v_s),(''bdv_2'', NTH 2 v_s ),     " ^
   683            "                              (''bdv_3'', NTH 3 v_s),(''bdv_3'', NTH 4 v_s )]     " ^
   684            "                             ''isolate_bdvs_4x4'' False))                 @@    " ^
   685            "      (Try (Rewrite_Set_Inst [(''bdv_1'', NTH 1 v_s),(''bdv_2'', NTH 2 v_s ),     " ^
   686            "                              (''bdv_3'', NTH 3 v_s),(''bdv_3'', NTH 4 v_s )]     " ^
   687            "                             ''simplify_System_parenthesized'' False))    @@    " ^
   688            "      (Try (Rewrite_Set ''order_system'' False)))                           e_s " ^
   689            "   in (SubProblem (''EqSystem'',[''LINEAR'',''system''],[''no_met''])                      " ^
   690            "                  [BOOL_LIST e__s, REAL_LIST v_s]))"*))]
   691 \<close>
   692 
   693 partial_function (tailrec) solve_system4 :: "bool list => real list => bool list"
   694   where
   695 "solve_system4 e_s v_s =
   696   (let e_1 = NTH 1 e_s;
   697        e_2 = Take (NTH 2 e_s);
   698        e_2 = ((Substitute [e_1]) @@
   699               (Try (Rewrite_Set_Inst [(''bdv_1'',NTH 1 v_s),(''bdv_2'',NTH 2 v_s),
   700                                       (''bdv_3'',NTH 3 v_s),(''bdv_4'',NTH 4 v_s)]
   701                                  ''simplify_System_parenthesized'' False)) @@
   702               (Try (Rewrite_Set_Inst [(''bdv_1'',NTH 1 v_s),(''bdv_2'',NTH 2 v_s),
   703                                       (''bdv_3'',NTH 3 v_s),(''bdv_4'',NTH 4 v_s)]
   704                                  ''isolate_bdvs'' False)) @@
   705               (Try (Rewrite_Set_Inst [(''bdv_1'',NTH 1 v_s),(''bdv_2'',NTH 2 v_s),
   706                                       (''bdv_3'',NTH 3 v_s),(''bdv_4'',NTH 4 v_s)]
   707                                  ''norm_Rational'' False))) e_2
   708     in [e_1, e_2, NTH 3 e_s, NTH 4 e_s])"
   709 setup \<open>KEStore_Elems.add_mets
   710     [Specify.prep_met thy "met_eqsys_topdown_4x4" [] Celem.e_metID
   711 	    (["EqSystem","top_down_substitution","4x4"],
   712 	      [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   713 	        ("#Where" , (*accepts missing variables up to diagonal form*)
   714             ["(NTH 1 (v_s::real list)) occurs_in (NTH 1 (e_s::bool list))",
   715               "(NTH 2 (v_s::real list)) occurs_in (NTH 2 (e_s::bool list))",
   716               "(NTH 3 (v_s::real list)) occurs_in (NTH 3 (e_s::bool list))",
   717               "(NTH 4 (v_s::real list)) occurs_in (NTH 4 (e_s::bool list))"]),
   718 	        ("#Find", ["solution ss'''"])],
   719 	    {rew_ord'="ord_simplify_System", rls' = Rule.Erls, calc = [], 
   720 	      srls = Rule.append_rls "srls_top_down_4x4" srls [], 
   721 	      prls = Rule.append_rls "prls_tri_4x4_lin_sys" prls_triangular
   722 			      [Rule.Calc ("Atools.occurs'_in",eval_occurs_in "")], 
   723 	      crls = Rule.Erls, errpats = [], nrls = Rule.Erls},
   724 	    (*FIXXXXME.WN060916: this script works ONLY for exp 7.79 @@@@@@@@@@@@@@@@@@@@*)
   725 	    @{thm solve_system4.simps}
   726 	    (*"Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   727         "  (let e_1 = NTH 1 e_s;                                                    " ^
   728         "       e_2 = Take (NTH 2 e_s);                                             " ^
   729         "       e_2 = ((Substitute [e_1]) @@                                         " ^
   730         "              (Try (Rewrite_Set_Inst [(''bdv_1'',NTH 1 v_s),(''bdv_2'',NTH 2 v_s)," ^
   731         "                                      (''bdv_3'',NTH 3 v_s),(''bdv_4'',NTH 4 v_s)]" ^
   732         "                                 ''simplify_System_parenthesized'' False)) @@   " ^
   733         "              (Try (Rewrite_Set_Inst [(''bdv_1'',NTH 1 v_s),(''bdv_2'',NTH 2 v_s)," ^
   734         "                                      (''bdv_3'',NTH 3 v_s),(''bdv_4'',NTH 4 v_s)]" ^
   735         "                                 ''isolate_bdvs'' False))                  @@   " ^
   736         "              (Try (Rewrite_Set_Inst [(''bdv_1'',NTH 1 v_s),(''bdv_2'',NTH 2 v_s)," ^
   737         "                                      (''bdv_3'',NTH 3 v_s),(''bdv_4'',NTH 4 v_s)]" ^
   738         "                                 ''norm_Rational'' False)))             e_2     " ^
   739         "    in [e_1, e_2, NTH 3 e_s, NTH 4 e_s])"*))]
   740 \<close>
   741 
   742 end