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