src/Tools/isac/Knowledge/EqSystem.thy
author Mathias Lehnfeld <s1210629013@students.fh-hagenberg.at>
Mon, 27 Jan 2014 21:49:27 +0100
changeset 55359 73dc85c025ab
parent 55339 cccd24e959ba
child 55363 d78bc1342183
permissions -rw-r--r--
cleanup, naming: 'KEStore_Elems' in Tests now 'Test_KEStore_Elems', 'store_pbts' now 'add_pbts'
     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,normalize,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 {*
    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 (isalist2list vs) (isalist2list all) t
    80     then SOME ((term2str p) ^ " = True",
    81 	       Trueprop $ (mk_equality (p, @{term True})))
    82     else SOME ((term2str p) ^ " = False",
    83 	       Trueprop $ (mk_equality (p, @{term False})))
    84   | eval_occur_exactly_in _ _ _ _ = NONE;
    85 *}
    86 setup {* KEStore_Elems.add_calcs
    87   [("occur_exactly_in",
    88 	    ("EqSystem.occur'_exactly'_in",
    89 	      eval_occur_exactly_in "#eval_occur_exactly_in_"))] *}
    90 ML {*
    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 * ((str2int 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= \"" ^ term_to_string''' thy f ^ "\" @ \"[" ^
   135            commas (map (term_to_string''' thy) ts) ^ "]\"");
   136          val _ = tracing ("u= g@us= \"" ^ term_to_string''' thy g ^ "\" @ \"[" ^
   137            commas (map (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 (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 rew_ord' := overwritel (!rew_ord',
   170 [("ord_simplify_System", ord_simplify_System false thy)
   171  ]);
   172 *}
   173 ML {*
   174 (** rulesets **)
   175 
   176 (*.adapted from 'order_add_mult_in' by just replacing the rew_ord.*)
   177 val order_add_mult_System = 
   178   Rls{id = "order_add_mult_System", preconds = [], 
   179       rew_ord = ("ord_simplify_System",
   180 		 ord_simplify_System false @{theory "Integrate"}),
   181       erls = e_rls,srls = Erls, calc = [], errpatts = [],
   182       rules = [Thm ("mult_commute",num_str @{thm mult_commute}),
   183 	       (* z * w = w * z *)
   184 	       Thm ("real_mult_left_commute",num_str @{thm real_mult_left_commute}),
   185 	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
   186 	       Thm ("mult_assoc",num_str @{thm mult_assoc}),		
   187 	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
   188 	       Thm ("add_commute",num_str @{thm add_commute}),	
   189 	       (*z + w = w + z*)
   190 	       Thm ("add_left_commute",num_str @{thm add_left_commute}),
   191 	       (*x + (y + z) = y + (x + z)*)
   192 	       Thm ("add_assoc",num_str @{thm add_assoc})	               
   193 	       (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
   194 	       ], 
   195       scr = EmptyScr}:rls;
   196 *}
   197 ML {*
   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   Rls {id = "norm_System_noadd_fractions", preconds = [], 
   203        rew_ord = ("dummy_ord",dummy_ord), 
   204        erls = norm_rat_erls, srls = Erls, calc = [], errpatts = [],
   205        rules = [(*sequence given by operator precedence*)
   206 		Rls_ discard_minus,
   207 		Rls_ powers,
   208 		Rls_ rat_mult_divide,
   209 		Rls_ expand,
   210 		Rls_ reduce_0_1_2,
   211 		Rls_ (*order_add_mult #1*) order_add_mult_System,
   212 		Rls_ collect_numerals,
   213 		(*Rls_ add_fractions_p, #2*)
   214 		Rls_ cancel_p
   215 		],
   216        scr = Prog ((term_of o the o (parse thy)) "empty_script")
   217        }:rls;
   218 *}
   219 ML {*
   220 (*.adapted from 'norm_Rational' by
   221   *1* using 'ord_simplify_System' in 'order_add_mult_System'.*)
   222 val norm_System = 
   223   Rls {id = "norm_System", preconds = [], 
   224        rew_ord = ("dummy_ord",dummy_ord), 
   225        erls = norm_rat_erls, srls = Erls, calc = [], errpatts = [],
   226        rules = [(*sequence given by operator precedence*)
   227 		Rls_ discard_minus,
   228 		Rls_ powers,
   229 		Rls_ rat_mult_divide,
   230 		Rls_ expand,
   231 		Rls_ reduce_0_1_2,
   232 		Rls_ (*order_add_mult *1*) order_add_mult_System,
   233 		Rls_ collect_numerals,
   234 		Rls_ add_fractions_p,
   235 		Rls_ cancel_p
   236 		],
   237        scr = Prog ((term_of o the o (parse thy)) "empty_script")
   238        }:rls;
   239 *}
   240 ML {*
   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   Seq {id = "simplify_System_parenthesized", preconds = []:term list, 
   252        rew_ord = ("dummy_ord", dummy_ord),
   253       erls = Atools_erls, srls = Erls, calc = [], errpatts = [],
   254       rules = [Thm ("distrib_right",num_str @{thm distrib_right}),
   255  	       (*"(?z1.0 + ?z2.0) * ?w = ?z1.0 * ?w + ?z2.0 * ?w"*)
   256 	       Thm ("add_divide_distrib",num_str @{thm add_divide_distrib}),
   257  	       (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)
   258 	       (*^^^^^ *0* ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^*)
   259 	       Rls_ norm_Rational_noadd_fractions(**2**),
   260 	       Rls_ (*order_add_mult_in*) norm_System_noadd_fractions (**1**),
   261 	       Thm ("sym_mult_assoc",
   262                      num_str (@{thm mult_assoc} RS @{thm sym}))
   263 	       (*Rls_ discard_parentheses *3**),
   264 	       Rls_ collect_bdv, (*from make_polynomial_in WN051031 welldone?*)
   265 	       Rls_ separate_bdv2,
   266 	       Calc ("Fields.inverse_class.divide"  ,eval_cancel "#divide_e")
   267 	       ],
   268       scr = EmptyScr}:rls;      
   269 *}
   270 ML {*
   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   Seq {id = "simplify_System", preconds = []:term list, 
   277        rew_ord = ("dummy_ord", dummy_ord),
   278       erls = Atools_erls, srls = Erls, calc = [], errpatts = [],
   279       rules = [Rls_ norm_Rational,
   280 	       Rls_ (*order_add_mult_in*) norm_System (**1**),
   281 	       Rls_ discard_parentheses,
   282 	       Rls_ collect_bdv, (*from make_polynomial_in WN051031 welldone?*)
   283 	       Rls_ separate_bdv2,
   284 	       Calc ("Fields.inverse_class.divide"  ,eval_cancel "#divide_e")
   285 	       ],
   286       scr = EmptyScr}:rls;      
   287 (*
   288 val simplify_System = 
   289     append_rls "simplify_System" simplify_System_parenthesized
   290 	       [Thm ("sym_add_assoc",
   291                       num_str (@{thm add_assoc} RS @{thm sym}))];
   292 *)
   293 *}
   294 ML {*
   295 val isolate_bdvs = 
   296     Rls {id="isolate_bdvs", preconds = [], 
   297 	 rew_ord = ("e_rew_ord", e_rew_ord), 
   298 	 erls = append_rls "erls_isolate_bdvs" e_rls 
   299 			   [(Calc ("EqSystem.occur'_exactly'_in", 
   300 				   eval_occur_exactly_in 
   301 				       "#eval_occur_exactly_in_"))
   302 			    ], 
   303 			   srls = Erls, calc = [], errpatts = [],
   304 	      rules = 
   305              [Thm ("commute_0_equality", num_str @{thm commute_0_equality}),
   306 	      Thm ("separate_bdvs_add", num_str @{thm separate_bdvs_add}),
   307 	      Thm ("separate_bdvs_mult", num_str @{thm separate_bdvs_mult})],
   308 	      scr = EmptyScr};
   309 *}
   310 ML {*
   311 val isolate_bdvs_4x4 = 
   312     Rls {id="isolate_bdvs_4x4", preconds = [], 
   313 	 rew_ord = ("e_rew_ord", e_rew_ord), 
   314 	 erls = append_rls 
   315 		    "erls_isolate_bdvs_4x4" e_rls 
   316 		    [Calc ("EqSystem.occur'_exactly'_in", 
   317 			   eval_occur_exactly_in "#eval_occur_exactly_in_"),
   318 		     Calc ("Atools.ident",eval_ident "#ident_"),
   319 		     Calc ("Atools.some'_occur'_in", 
   320 			   eval_some_occur_in "#some_occur_in_"),
   321                      Thm ("not_true",num_str @{thm not_true}),
   322 		     Thm ("not_false",num_str @{thm not_false})
   323 			    ], 
   324 	 srls = Erls, calc = [], errpatts = [],
   325 	 rules = [Thm ("commute_0_equality", num_str @{thm commute_0_equality}),
   326 		  Thm ("separate_bdvs0", num_str @{thm separate_bdvs0}),
   327 		  Thm ("separate_bdvs_add1", num_str @{thm separate_bdvs_add1}),
   328 		  Thm ("separate_bdvs_add1", num_str @{thm separate_bdvs_add2}),
   329 		  Thm ("separate_bdvs_mult", num_str @{thm separate_bdvs_mult})
   330                  ], scr = EmptyScr};
   331 
   332 *}
   333 ML {*
   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     Rls {id="order_system", preconds = [], 
   339 	 rew_ord = ("ord_simplify_System", 
   340 		    ord_simplify_System false thy), 
   341 	 erls = Erls, srls = Erls, calc = [], errpatts = [],
   342 	 rules = [Thm ("order_system_NxN", num_str @{thm order_system_NxN})
   343 		  ],
   344 	 scr = EmptyScr};
   345 
   346 val prls_triangular = 
   347     Rls {id="prls_triangular", preconds = [], 
   348 	 rew_ord = ("e_rew_ord", e_rew_ord), 
   349 	 erls = Rls {id="erls_prls_triangular", preconds = [], 
   350 		     rew_ord = ("e_rew_ord", e_rew_ord), 
   351 		     erls = Erls, srls = Erls, calc = [], errpatts = [],
   352 		     rules = [(*for precond NTH_CONS ...*)
   353 			      Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   354 			      Calc ("Groups.plus_class.plus", eval_binop "#add_")
   355 			      (*immediately repeated rewrite pushes
   356 					    '+' into precondition !*)
   357 			      ],
   358 		     scr = EmptyScr}, 
   359 	 srls = Erls, calc = [], errpatts = [],
   360 	 rules = [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   361 		  Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   362 		  Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
   363 		  Thm ("tl_Cons",num_str @{thm tl_Cons}),
   364 		  Thm ("tl_Nil",num_str @{thm tl_Nil}),
   365 		  Calc ("EqSystem.occur'_exactly'_in", 
   366 			eval_occur_exactly_in 
   367 			    "#eval_occur_exactly_in_")
   368 		  ],
   369 	 scr = EmptyScr};
   370 *}
   371 ML {*
   372 
   373 (*WN060914 quickly created for 4x4; 
   374  more similarity to prls_triangular desirable*)
   375 val prls_triangular4 = 
   376     Rls {id="prls_triangular4", preconds = [], 
   377 	 rew_ord = ("e_rew_ord", e_rew_ord), 
   378 	 erls = Rls {id="erls_prls_triangular4", preconds = [], 
   379 		     rew_ord = ("e_rew_ord", e_rew_ord), 
   380 		     erls = Erls, srls = Erls, calc = [], errpatts = [],
   381 		     rules = [(*for precond NTH_CONS ...*)
   382 			      Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   383 			      Calc ("Groups.plus_class.plus", eval_binop "#add_")
   384 			      (*immediately repeated rewrite pushes
   385 					    '+' into precondition !*)
   386 			      ],
   387 		     scr = EmptyScr}, 
   388 	 srls = Erls, calc = [], errpatts = [],
   389 	 rules = [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   390 		  Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   391 		  Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
   392 		  Thm ("tl_Cons",num_str @{thm tl_Cons}),
   393 		  Thm ("tl_Nil",num_str @{thm tl_Nil}),
   394 		  Calc ("EqSystem.occur'_exactly'_in", 
   395 			eval_occur_exactly_in 
   396 			    "#eval_occur_exactly_in_")
   397 		  ],
   398 	 scr = EmptyScr};
   399 *}
   400 
   401 setup {* 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))] *}
   412 
   413 ML {*
   414 (** problems **)
   415 
   416 store_pbt
   417  (prep_pbt thy "pbl_equsys" [] e_pblID
   418  (["system"],
   419   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   420    ("#Find"  ,["solution ss'''"](*''' is copy-named*))
   421   ],
   422   append_rls "e_rls" e_rls [(*for preds in where_*)], 
   423   SOME "solveSystem e_s v_s", 
   424   []));
   425 store_pbt
   426  (prep_pbt thy "pbl_equsys_lin" [] e_pblID
   427  (["LINEAR", "system"],
   428   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   429    (*TODO.WN050929 check linearity*)
   430    ("#Find"  ,["solution ss'''"])
   431   ],
   432   append_rls "e_rls" e_rls [(*for preds in where_*)], 
   433   SOME "solveSystem e_s v_s", 
   434   []));
   435 store_pbt
   436  (prep_pbt thy "pbl_equsys_lin_2x2" [] e_pblID
   437  (["2x2", "LINEAR", "system"],
   438   (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   439   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   440    ("#Where"  ,["LENGTH (e_s:: bool list) = 2", "LENGTH v_s = 2"]),
   441    ("#Find"  ,["solution ss'''"])
   442   ],
   443   append_rls "prls_2x2_linear_system" e_rls 
   444 			     [Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   445 			      Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   446 			      Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   447 			      Calc ("HOL.eq",eval_equal "#equal_")
   448 			      ], 
   449   SOME "solveSystem e_s v_s", 
   450   []));
   451 *}
   452 ML {*
   453 store_pbt
   454  (prep_pbt thy "pbl_equsys_lin_2x2_tri" [] e_pblID
   455  (["triangular", "2x2", "LINEAR", "system"],
   456   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   457    ("#Where"  ,
   458     ["(tl v_s) from v_s occur_exactly_in (NTH 1 (e_s::bool list))",
   459      "    v_s  from v_s occur_exactly_in (NTH 2 (e_s::bool list))"]),
   460    ("#Find"  ,["solution ss'''"])
   461   ],
   462   prls_triangular, 
   463   SOME "solveSystem e_s v_s", 
   464   [["EqSystem","top_down_substitution","2x2"]]));
   465 store_pbt
   466  (prep_pbt thy "pbl_equsys_lin_2x2_norm" [] e_pblID
   467  (["normalize", "2x2", "LINEAR", "system"],
   468   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   469    ("#Find"  ,["solution ss'''"])
   470   ],
   471   append_rls "e_rls" e_rls [(*for preds in where_*)], 
   472   SOME "solveSystem e_s v_s", 
   473   [["EqSystem","normalize","2x2"]]));
   474 store_pbt
   475  (prep_pbt thy "pbl_equsys_lin_3x3" [] e_pblID
   476  (["3x3", "LINEAR", "system"],
   477   (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   478   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   479    ("#Where"  ,["LENGTH (e_s:: bool list) = 3", "LENGTH v_s = 3"]),
   480    ("#Find"  ,["solution ss'''"])
   481   ],
   482   append_rls "prls_3x3_linear_system" e_rls 
   483 			     [Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   484 			      Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   485 			      Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   486 			      Calc ("HOL.eq",eval_equal "#equal_")
   487 			      ], 
   488   SOME "solveSystem e_s v_s", 
   489   []));
   490 store_pbt
   491  (prep_pbt thy "pbl_equsys_lin_4x4" [] e_pblID
   492  (["4x4", "LINEAR", "system"],
   493   (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   494   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   495    ("#Where"  ,["LENGTH (e_s:: bool list) = 4", "LENGTH v_s = 4"]),
   496    ("#Find"  ,["solution ss'''"])
   497   ],
   498   append_rls "prls_4x4_linear_system" e_rls 
   499 			     [Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   500 			      Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   501 			      Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   502 			      Calc ("HOL.eq",eval_equal "#equal_")
   503 			      ], 
   504   SOME "solveSystem e_s v_s", 
   505   []));
   506 *}
   507 ML {*
   508 store_pbt
   509  (prep_pbt thy "pbl_equsys_lin_4x4_tri" [] e_pblID
   510  (["triangular", "4x4", "LINEAR", "system"],
   511   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   512    ("#Where" , (*accepts missing variables up to diagional form*)
   513     ["(NTH 1 (v_s::real list)) occurs_in (NTH 1 (e_s::bool list))",
   514      "(NTH 2 (v_s::real list)) occurs_in (NTH 2 (e_s::bool list))",
   515      "(NTH 3 (v_s::real list)) occurs_in (NTH 3 (e_s::bool list))",
   516      "(NTH 4 (v_s::real list)) occurs_in (NTH 4 (e_s::bool list))"
   517      ]),
   518    ("#Find"  ,["solution ss'''"])
   519   ],
   520   append_rls "prls_tri_4x4_lin_sys" prls_triangular
   521 	     [Calc ("Atools.occurs'_in",eval_occurs_in "")], 
   522   SOME "solveSystem e_s v_s", 
   523   [["EqSystem","top_down_substitution","4x4"]]));
   524 *}
   525 ML {*
   526 store_pbt
   527  (prep_pbt thy "pbl_equsys_lin_4x4_norm" [] e_pblID
   528  (["normalize", "4x4", "LINEAR", "system"],
   529   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   530    (*LENGTH is checked 1 level above*)
   531    ("#Find"  ,["solution ss'''"])
   532   ],
   533   append_rls "e_rls" e_rls [(*for preds in where_*)], 
   534   SOME "solveSystem e_s v_s", 
   535   [["EqSystem","normalize","4x4"]]));
   536 
   537 
   538 (* show_ptyps();
   539    *)
   540 
   541 *}
   542 setup {* KEStore_Elems.add_pbts
   543   [(prep_pbt thy "pbl_equsys" [] e_pblID
   544       (["system"],
   545         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   546           ("#Find"  ,["solution ss'''"](*''' is copy-named*))],
   547         append_rls "e_rls" e_rls [(*for preds in where_*)], SOME "solveSystem e_s v_s", [])),
   548     (prep_pbt thy "pbl_equsys_lin" [] e_pblID
   549       (["LINEAR", "system"],
   550         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   551           (*TODO.WN050929 check linearity*)
   552           ("#Find"  ,["solution ss'''"])],
   553         append_rls "e_rls" e_rls [(*for preds in where_*)], SOME "solveSystem e_s v_s", [])),
   554     (prep_pbt thy "pbl_equsys_lin_2x2" [] e_pblID
   555       (["2x2", "LINEAR", "system"],
   556       (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   557         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   558           ("#Where"  ,["LENGTH (e_s:: bool list) = 2", "LENGTH v_s = 2"]),
   559           ("#Find"  ,["solution ss'''"])],
   560         append_rls "prls_2x2_linear_system" e_rls 
   561 			    [Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   562 			      Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   563 			      Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   564 			      Calc ("HOL.eq",eval_equal "#equal_")], 
   565         SOME "solveSystem e_s v_s", [])),
   566     (prep_pbt thy "pbl_equsys_lin_2x2_tri" [] e_pblID
   567       (["triangular", "2x2", "LINEAR", "system"],
   568         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   569           ("#Where",
   570             ["(tl v_s) from v_s occur_exactly_in (NTH 1 (e_s::bool list))",
   571               "    v_s  from v_s occur_exactly_in (NTH 2 (e_s::bool list))"]),
   572           ("#Find"  ,["solution ss'''"])],
   573         prls_triangular, SOME "solveSystem e_s v_s", [["EqSystem","top_down_substitution","2x2"]])),
   574     (prep_pbt thy "pbl_equsys_lin_2x2_norm" [] e_pblID
   575       (["normalize", "2x2", "LINEAR", "system"],
   576         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   577           ("#Find"  ,["solution ss'''"])],
   578       append_rls "e_rls" e_rls [(*for preds in where_*)], 
   579       SOME "solveSystem e_s v_s", 
   580       [["EqSystem","normalize","2x2"]])),
   581     (prep_pbt thy "pbl_equsys_lin_3x3" [] e_pblID
   582       (["3x3", "LINEAR", "system"],
   583         (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   584         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   585           ("#Where"  ,["LENGTH (e_s:: bool list) = 3", "LENGTH v_s = 3"]),
   586           ("#Find"  ,["solution ss'''"])],
   587         append_rls "prls_3x3_linear_system" e_rls 
   588 			    [Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   589 			      Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   590 			      Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   591 			      Calc ("HOL.eq",eval_equal "#equal_")],
   592         SOME "solveSystem e_s v_s", [])),
   593     (prep_pbt thy "pbl_equsys_lin_4x4" [] e_pblID
   594       (["4x4", "LINEAR", "system"],
   595         (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   596         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   597           ("#Where"  ,["LENGTH (e_s:: bool list) = 4", "LENGTH v_s = 4"]),
   598           ("#Find"  ,["solution ss'''"])],
   599         append_rls "prls_4x4_linear_system" e_rls 
   600 			    [Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   601 			      Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   602 			      Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   603 			      Calc ("HOL.eq",eval_equal "#equal_")],
   604         SOME "solveSystem e_s v_s", [])),
   605     (prep_pbt thy "pbl_equsys_lin_4x4_tri" [] e_pblID
   606       (["triangular", "4x4", "LINEAR", "system"],
   607         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   608           ("#Where" , (*accepts missing variables up to diagional form*)
   609             ["(NTH 1 (v_s::real list)) occurs_in (NTH 1 (e_s::bool list))",
   610               "(NTH 2 (v_s::real list)) occurs_in (NTH 2 (e_s::bool list))",
   611               "(NTH 3 (v_s::real list)) occurs_in (NTH 3 (e_s::bool list))",
   612               "(NTH 4 (v_s::real list)) occurs_in (NTH 4 (e_s::bool list))"]),
   613           ("#Find"  ,["solution ss'''"])],
   614       append_rls "prls_tri_4x4_lin_sys" prls_triangular
   615 	      [Calc ("Atools.occurs'_in",eval_occurs_in "")], 
   616       SOME "solveSystem e_s v_s", 
   617       [["EqSystem","top_down_substitution","4x4"]])),
   618     (prep_pbt thy "pbl_equsys_lin_4x4_norm" [] e_pblID
   619       (["normalize", "4x4", "LINEAR", "system"],
   620         [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   621           (*LENGTH is checked 1 level above*)
   622           ("#Find"  ,["solution ss'''"])],
   623         append_rls "e_rls" e_rls [(*for preds in where_*)], 
   624         SOME "solveSystem e_s v_s", 
   625         [["EqSystem","normalize","4x4"]]))] *}
   626 ML {*
   627 (** methods **)
   628 
   629 store_met
   630     (prep_met thy "met_eqsys" [] e_metID
   631 	      (["EqSystem"],
   632 	       [],
   633 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   634 		srls = Erls, prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   635 	       "empty_script"
   636 	       ));
   637 store_met
   638     (prep_met thy "met_eqsys_topdown" [] e_metID
   639 	      (["EqSystem","top_down_substitution"],
   640 	       [],
   641 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   642 		srls = Erls, prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   643 	       "empty_script"
   644 	       ));
   645 *}
   646 ML {*
   647 store_met
   648     (prep_met thy "met_eqsys_topdown_2x2" [] e_metID
   649 	 (["EqSystem","top_down_substitution","2x2"],
   650 	  [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   651 	   ("#Where"  ,
   652 	    ["(tl v_s) from v_s occur_exactly_in (NTH 1 (e_s::bool list))",
   653 	     "    v_s  from v_s occur_exactly_in (NTH 2 (e_s::bool list))"]),
   654 	   ("#Find"  ,["solution ss'''"])
   655 	   ],
   656 	  {rew_ord'="ord_simplify_System", rls' = Erls, calc = [], 
   657 	   srls = append_rls "srls_top_down_2x2" e_rls
   658 				  [Thm ("hd_thm",num_str @{thm hd_thm}),
   659 				   Thm ("tl_Cons",num_str @{thm tl_Cons}),
   660 				   Thm ("tl_Nil",num_str @{thm tl_Nil})
   661 				   ], 
   662 	   prls = prls_triangular, crls = Erls, errpats = [], nrls = Erls},
   663 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   664 "  (let e_1 = Take (hd e_s);                                                " ^
   665 "       e_1 = ((Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   666 "                                  isolate_bdvs False))     @@               " ^
   667 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   668 "                                  simplify_System False))) e_1;            " ^
   669 "       e_2 = Take (hd (tl e_s));                                           " ^
   670 "       e_2 = ((Substitute [e_1]) @@                                       " ^
   671 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   672 "                                  simplify_System_parenthesized False)) @@  " ^
   673 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   674 "                                  isolate_bdvs False))     @@               " ^
   675 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   676 "                                  simplify_System False))) e_2;            " ^
   677 "       e__s = Take [e_1, e_2]                                             " ^
   678 "   in (Try (Rewrite_Set order_system False)) e__s)"
   679 (*---------------------------------------------------------------------------
   680   this script does NOT separate the equations as above, 
   681   but it does not yet work due to preliminary script-interpreter,
   682   see eqsystem.sml 'script [EqSystem,top_down_substitution,2x2] Vers.2'
   683 
   684 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =         " ^
   685 "  (let e__s = Take e_s;                                              " ^
   686 "       e_1 = hd e__s;                                               " ^
   687 "       e_2 = hd (tl e__s);                                          " ^
   688 "       e__s = [e_1, Substitute [e_1] e_2]                         " ^
   689 "   in ((Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   690 "                                  simplify_System_parenthesized False)) @@ " ^
   691 "       (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))] " ^
   692 "                              isolate_bdvs False))              @@   " ^
   693 "       (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   694 "                                  simplify_System False))) e__s)"
   695 ---------------------------------------------------------------------------*)
   696 	  ));
   697 *}
   698 ML {*
   699 store_met
   700     (prep_met thy "met_eqsys_norm" [] e_metID
   701 	      (["EqSystem","normalize"],
   702 	       [],
   703 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   704 		srls = Erls, prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   705 	       "empty_script"
   706 	       ));
   707 *}
   708 ML {*
   709 store_met
   710     (prep_met thy "met_eqsys_norm_2x2" [] e_metID
   711 	      (["EqSystem","normalize","2x2"],
   712 	       [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   713 		("#Find"  ,["solution ss'''"])],
   714 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   715 		srls = append_rls "srls_normalize_2x2" e_rls
   716 				  [Thm ("hd_thm",num_str @{thm hd_thm}),
   717 				   Thm ("tl_Cons",num_str @{thm tl_Cons}),
   718 				   Thm ("tl_Nil",num_str @{thm tl_Nil})
   719 				   ], 
   720 		prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   721 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   722 "  (let e__s = ((Try (Rewrite_Set norm_Rational False)) @@                   " ^
   723 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   724 "                                  simplify_System_parenthesized False)) @@  " ^
   725 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   726 "                                                    isolate_bdvs False)) @@ " ^
   727 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   728 "                                  simplify_System_parenthesized False)) @@  " ^
   729 "               (Try (Rewrite_Set order_system False))) e_s                  " ^
   730 "   in (SubProblem (EqSystem',[LINEAR,system],[no_met])                      " ^
   731 "                  [BOOL_LIST e__s, REAL_LIST v_s]))"
   732 	       ));
   733 
   734 *}
   735 ML {*
   736 (*this is for NTH only*)
   737 val srls = Rls {id="srls_normalize_4x4", 
   738 		preconds = [], 
   739 		rew_ord = ("termlessI",termlessI), 
   740 		erls = append_rls "erls_in_srls_IntegrierenUnd.." e_rls
   741 				  [(*for asm in NTH_CONS ...*)
   742 				   Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   743 				   (*2nd NTH_CONS pushes n+-1 into asms*)
   744 				   Calc("Groups.plus_class.plus", eval_binop "#add_")
   745 				   ], 
   746 		srls = Erls, calc = [], errpatts = [],
   747 		rules = [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   748 			 Calc("Groups.plus_class.plus", eval_binop "#add_"),
   749 			 Thm ("NTH_NIL",num_str @{thm NTH_NIL})],
   750 		scr = EmptyScr};
   751 
   752 store_met
   753     (prep_met thy "met_eqsys_norm_4x4" [] e_metID
   754 	      (["EqSystem","normalize","4x4"],
   755 	       [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   756 		("#Find"  ,["solution ss'''"])],
   757 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   758 		srls = append_rls "srls_normalize_4x4" srls
   759 				  [Thm ("hd_thm",num_str @{thm hd_thm}),
   760 				   Thm ("tl_Cons",num_str @{thm tl_Cons}),
   761 				   Thm ("tl_Nil",num_str @{thm tl_Nil})
   762 				   ], 
   763 		prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   764 (*STOPPED.WN06? met ["EqSystem","normalize","4x4"] @@@@@@@@@@@@@@@@@@@@@@@@@@@*)
   765 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   766 "  (let e__s =                                                               " ^
   767 "     ((Try (Rewrite_Set norm_Rational False)) @@                            " ^
   768 "      (Repeat (Rewrite commute_0_equality False)) @@                        " ^
   769 "      (Try (Rewrite_Set_Inst [(bdv_1, NTH 1 v_s),(bdv_2, NTH 2 v_s ),     " ^
   770 "                              (bdv_3, NTH 3 v_s),(bdv_3, NTH 4 v_s )]     " ^
   771 "                             simplify_System_parenthesized False))    @@    " ^
   772 "      (Try (Rewrite_Set_Inst [(bdv_1, NTH 1 v_s),(bdv_2, NTH 2 v_s ),     " ^
   773 "                              (bdv_3, NTH 3 v_s),(bdv_3, NTH 4 v_s )]     " ^
   774 "                             isolate_bdvs_4x4 False))                 @@    " ^
   775 "      (Try (Rewrite_Set_Inst [(bdv_1, NTH 1 v_s),(bdv_2, NTH 2 v_s ),     " ^
   776 "                              (bdv_3, NTH 3 v_s),(bdv_3, NTH 4 v_s )]     " ^
   777 "                             simplify_System_parenthesized False))    @@    " ^
   778 "      (Try (Rewrite_Set order_system False)))                           e_s " ^
   779 "   in (SubProblem (EqSystem',[LINEAR,system],[no_met])                      " ^
   780 "                  [BOOL_LIST e__s, REAL_LIST v_s]))"
   781 ));
   782 
   783 store_met
   784 (prep_met thy "met_eqsys_topdown_4x4" [] e_metID
   785 	  (["EqSystem","top_down_substitution","4x4"],
   786 	   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   787 	    ("#Where" , (*accepts missing variables up to diagonal form*)
   788 	     ["(NTH 1 (v_s::real list)) occurs_in (NTH 1 (e_s::bool list))",
   789 	      "(NTH 2 (v_s::real list)) occurs_in (NTH 2 (e_s::bool list))",
   790 	      "(NTH 3 (v_s::real list)) occurs_in (NTH 3 (e_s::bool list))",
   791 	      "(NTH 4 (v_s::real list)) occurs_in (NTH 4 (e_s::bool list))"
   792 	      ]),
   793 	    ("#Find"  ,["solution ss'''"])
   794 	    ],
   795 	   {rew_ord'="ord_simplify_System", rls' = Erls, calc = [], 
   796 	    srls = append_rls "srls_top_down_4x4" srls [], 
   797 	    prls = append_rls "prls_tri_4x4_lin_sys" prls_triangular
   798 			      [Calc ("Atools.occurs'_in",eval_occurs_in "")], 
   799 	    crls = Erls, errpats = [], nrls = Erls},
   800 (*FIXXXXME.WN060916: this script works ONLY for exp 7.79 @@@@@@@@@@@@@@@@@@@@*)
   801 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   802 "  (let e_1 = NTH 1 e_s;                                                    " ^
   803 "       e_2 = Take (NTH 2 e_s);                                             " ^
   804 "       e_2 = ((Substitute [e_1]) @@                                         " ^
   805 "              (Try (Rewrite_Set_Inst [(bdv_1,NTH 1 v_s),(bdv_2,NTH 2 v_s)," ^
   806 "                                      (bdv_3,NTH 3 v_s),(bdv_4,NTH 4 v_s)]" ^
   807 "                                 simplify_System_parenthesized False)) @@   " ^
   808 "              (Try (Rewrite_Set_Inst [(bdv_1,NTH 1 v_s),(bdv_2,NTH 2 v_s)," ^
   809 "                                      (bdv_3,NTH 3 v_s),(bdv_4,NTH 4 v_s)]" ^
   810 "                                 isolate_bdvs False))                  @@   " ^
   811 "              (Try (Rewrite_Set_Inst [(bdv_1,NTH 1 v_s),(bdv_2,NTH 2 v_s)," ^
   812 "                                      (bdv_3,NTH 3 v_s),(bdv_4,NTH 4 v_s)]" ^
   813 "                                 norm_Rational False)))             e_2     " ^
   814 "    in [e_1, e_2, NTH 3 e_s, NTH 4 e_s])"
   815 ));
   816 *}
   817 
   818 end