src/Tools/isac/Knowledge/EqSystem.thy
author Walther Neuper <neuper@ist.tugraz.at>
Wed, 08 Sep 2010 16:47:22 +0200
branchisac-update-Isa09-2
changeset 37991 028442673981
parent 37984 972a73d7c50b
child 37995 fac82f29f143
permissions -rw-r--r--
tuned src + test

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