src/Tools/isac/Knowledge/EqSystem.thy
author Walther Neuper <neuper@ist.tugraz.at>
Mon, 22 Jul 2013 13:52:18 +0200
changeset 52070 77138c64f4f6
parent 52062 b3f18f0d55d9
child 52125 6f1d3415dc68
permissions -rw-r--r--
--- Test_Isac.thy runs all tests

NOTEs:
(1) inserted many "trace_rewrite := false" ...
(2) ...nevertheless needs manual interaction at "Tracing paused. Stop, or continue with next 100, 1000, 10000 messages?"
(3) state of tests see respective section in Test_Isac.thy
     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 axioms(*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 axioms(*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 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 Symbol.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 Symbol.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 => Term_Ord.typ_ord (T, U) | ord => ord)
   132   | term_ord' pr thy (t, u) =
   133     (if pr
   134      then 
   135        let
   136          val (f, ts) = strip_comb t and (g, us) = strip_comb u;
   137          val _ = tracing ("t= f@ts= \"" ^ term_to_string''' thy f ^ "\" @ \"[" ^
   138            commas (map (term_to_string''' thy) ts) ^ "]\"");
   139          val _ = tracing ("u= g@us= \"" ^ term_to_string''' thy g ^ "\" @ \"[" ^
   140            commas (map (term_to_string''' thy) us) ^ "]\"");
   141          val _ = tracing ("size_of_term(t,u)= (" ^ string_of_int (size_of_term' t) ^ ", " ^
   142            string_of_int (size_of_term' u) ^ ")");
   143          val _ = tracing ("hd_ord(f,g)      = " ^ ((pr_ord o hd_ord) (f,g)));
   144          val _ = tracing ("terms_ord (ts,us) = " ^(pr_ord o terms_ord str false) (ts,us));
   145          val _=tracing("-------");
   146        in () end
   147      else ();
   148     case int_ord (size_of_term' t, size_of_term' u) of
   149       EQUAL =>
   150         let val (f, ts) = strip_comb t and (g, us) = strip_comb u 
   151         in (case hd_ord (f, g) of 
   152               EQUAL => (terms_ord str pr) (ts, us) 
   153             | ord => ord)
   154         end
   155 	 | ord => ord)
   156 and hd_ord (f, g) =                                        (* ~ term.ML *)
   157   prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord (dest_hd' f, dest_hd' g)
   158 and terms_ord str pr (ts, us) = list_ord (term_ord' pr (assoc_thy "Isac"))(ts, us);
   159 (**)
   160 in
   161 (**)
   162 (*WN0510 for preliminary use in eval_order_system, see case-study mat-eng.tex
   163 fun ord_simplify_System_rev (pr:bool) thy subst tu = 
   164     (term_ord' pr thy (Library.swap tu) = LESS);*)
   165 
   166 (*for the rls's*)
   167 fun ord_simplify_System (pr:bool) thy subst tu = 
   168     (term_ord' pr thy tu = LESS);
   169 (**)
   170 end;
   171 (**)
   172 rew_ord' := overwritel (!rew_ord',
   173 [("ord_simplify_System", ord_simplify_System false thy)
   174  ]);
   175 *}
   176 ML {*
   177 (** rulesets **)
   178 
   179 (*.adapted from 'order_add_mult_in' by just replacing the rew_ord.*)
   180 val order_add_mult_System = 
   181   Rls{id = "order_add_mult_System", preconds = [], 
   182       rew_ord = ("ord_simplify_System",
   183 		 ord_simplify_System false @{theory "Integrate"}),
   184       erls = e_rls,srls = Erls, calc = [], errpatts = [],
   185       rules = [Thm ("mult_commute",num_str @{thm mult_commute}),
   186 	       (* z * w = w * z *)
   187 	       Thm ("real_mult_left_commute",num_str @{thm real_mult_left_commute}),
   188 	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
   189 	       Thm ("mult_assoc",num_str @{thm mult_assoc}),		
   190 	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
   191 	       Thm ("add_commute",num_str @{thm add_commute}),	
   192 	       (*z + w = w + z*)
   193 	       Thm ("add_left_commute",num_str @{thm add_left_commute}),
   194 	       (*x + (y + z) = y + (x + z)*)
   195 	       Thm ("add_assoc",num_str @{thm add_assoc})	               
   196 	       (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
   197 	       ], 
   198       scr = EmptyScr}:rls;
   199 *}
   200 ML {*
   201 (*.adapted from 'norm_Rational' by
   202   #1 using 'ord_simplify_System' in 'order_add_mult_System'
   203   #2 NOT using common_nominator_p                          .*)
   204 val norm_System_noadd_fractions = 
   205   Rls {id = "norm_System_noadd_fractions", preconds = [], 
   206        rew_ord = ("dummy_ord",dummy_ord), 
   207        erls = norm_rat_erls, srls = Erls, calc = [], errpatts = [],
   208        rules = [(*sequence given by operator precedence*)
   209 		Rls_ discard_minus,
   210 		Rls_ powers,
   211 		Rls_ rat_mult_divide,
   212 		Rls_ expand,
   213 		Rls_ reduce_0_1_2,
   214 		Rls_ (*order_add_mult #1*) order_add_mult_System,
   215 		Rls_ collect_numerals,
   216 		(*Rls_ add_fractions_p, #2*)
   217 		Rls_ cancel_p
   218 		],
   219        scr = Prog ((term_of o the o (parse thy)) "empty_script")
   220        }:rls;
   221 *}
   222 ML {*
   223 (*.adapted from 'norm_Rational' by
   224   *1* using 'ord_simplify_System' in 'order_add_mult_System'.*)
   225 val norm_System = 
   226   Rls {id = "norm_System", preconds = [], 
   227        rew_ord = ("dummy_ord",dummy_ord), 
   228        erls = norm_rat_erls, srls = Erls, calc = [], errpatts = [],
   229        rules = [(*sequence given by operator precedence*)
   230 		Rls_ discard_minus,
   231 		Rls_ powers,
   232 		Rls_ rat_mult_divide,
   233 		Rls_ expand,
   234 		Rls_ reduce_0_1_2,
   235 		Rls_ (*order_add_mult *1*) order_add_mult_System,
   236 		Rls_ collect_numerals,
   237 		Rls_ add_fractions_p,
   238 		Rls_ cancel_p
   239 		],
   240        scr = Prog ((term_of o the o (parse thy)) "empty_script")
   241        }:rls;
   242 *}
   243 ML {*
   244 (*.simplify an equational system BEFORE solving it such that parentheses are
   245    ( ((u0*v0)*w0) + ( ((u1*v1)*w1) * c + ... +((u4*v4)*w4) * c_4 ) )
   246 ATTENTION: works ONLY for bound variables c, c_1, c_2, c_3, c_4 :ATTENTION
   247    This is a copy from 'make_ratpoly_in' with respective reductions:
   248    *0* expand the term, ie. distribute * and / over +
   249    *1* ord_simplify_System instead of termlessI
   250    *2* no add_fractions_p (= common_nominator_p_rls !)
   251    *3* discard_parentheses only for (.*(.*.))
   252    analoguous to simplify_Integral                                       .*)
   253 val simplify_System_parenthesized = 
   254   Seq {id = "simplify_System_parenthesized", preconds = []:term list, 
   255        rew_ord = ("dummy_ord", dummy_ord),
   256       erls = Atools_erls, srls = Erls, calc = [], errpatts = [],
   257       rules = [Thm ("distrib_right",num_str @{thm distrib_right}),
   258  	       (*"(?z1.0 + ?z2.0) * ?w = ?z1.0 * ?w + ?z2.0 * ?w"*)
   259 	       Thm ("add_divide_distrib",num_str @{thm add_divide_distrib}),
   260  	       (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)
   261 	       (*^^^^^ *0* ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^*)
   262 	       Rls_ norm_Rational_noadd_fractions(**2**),
   263 	       Rls_ (*order_add_mult_in*) norm_System_noadd_fractions (**1**),
   264 	       Thm ("sym_mult_assoc",
   265                      num_str (@{thm mult_assoc} RS @{thm sym}))
   266 	       (*Rls_ discard_parentheses *3**),
   267 	       Rls_ collect_bdv, (*from make_polynomial_in WN051031 welldone?*)
   268 	       Rls_ separate_bdv2,
   269 	       Calc ("Fields.inverse_class.divide"  ,eval_cancel "#divide_e")
   270 	       ],
   271       scr = EmptyScr}:rls;      
   272 *}
   273 ML {*
   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 = [], errpatts = [],
   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 ("Fields.inverse_class.divide"  ,eval_cancel "#divide_e")
   288 	       ],
   289       scr = EmptyScr}:rls;      
   290 (*
   291 val simplify_System = 
   292     append_rls "simplify_System" simplify_System_parenthesized
   293 	       [Thm ("sym_add_assoc",
   294                       num_str (@{thm add_assoc} RS @{thm sym}))];
   295 *)
   296 *}
   297 ML {*
   298 val isolate_bdvs = 
   299     Rls {id="isolate_bdvs", preconds = [], 
   300 	 rew_ord = ("e_rew_ord", e_rew_ord), 
   301 	 erls = append_rls "erls_isolate_bdvs" e_rls 
   302 			   [(Calc ("EqSystem.occur'_exactly'_in", 
   303 				   eval_occur_exactly_in 
   304 				       "#eval_occur_exactly_in_"))
   305 			    ], 
   306 			   srls = Erls, calc = [], errpatts = [],
   307 	      rules = 
   308              [Thm ("commute_0_equality", num_str @{thm commute_0_equality}),
   309 	      Thm ("separate_bdvs_add", num_str @{thm separate_bdvs_add}),
   310 	      Thm ("separate_bdvs_mult", num_str @{thm separate_bdvs_mult})],
   311 	      scr = EmptyScr};
   312 *}
   313 ML {*
   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 = [], errpatts = [],
   328 	 rules = [Thm ("commute_0_equality", num_str @{thm commute_0_equality}),
   329 		  Thm ("separate_bdvs0", num_str @{thm separate_bdvs0}),
   330 		  Thm ("separate_bdvs_add1", num_str @{thm separate_bdvs_add1}),
   331 		  Thm ("separate_bdvs_add1", num_str @{thm separate_bdvs_add2}),
   332 		  Thm ("separate_bdvs_mult", num_str @{thm separate_bdvs_mult})
   333                  ], scr = EmptyScr};
   334 
   335 *}
   336 ML {*
   337 
   338 (*.order the equations in a system such, that a triangular system (if any)
   339    appears as [..c_4 = .., ..., ..., ..c_1 + ..c_2 + ..c_3 ..c_4 = ..].*)
   340 val order_system = 
   341     Rls {id="order_system", preconds = [], 
   342 	 rew_ord = ("ord_simplify_System", 
   343 		    ord_simplify_System false thy), 
   344 	 erls = Erls, srls = Erls, calc = [], errpatts = [],
   345 	 rules = [Thm ("order_system_NxN", num_str @{thm order_system_NxN})
   346 		  ],
   347 	 scr = EmptyScr};
   348 
   349 val prls_triangular = 
   350     Rls {id="prls_triangular", preconds = [], 
   351 	 rew_ord = ("e_rew_ord", e_rew_ord), 
   352 	 erls = Rls {id="erls_prls_triangular", preconds = [], 
   353 		     rew_ord = ("e_rew_ord", e_rew_ord), 
   354 		     erls = Erls, srls = Erls, calc = [], errpatts = [],
   355 		     rules = [(*for precond NTH_CONS ...*)
   356 			      Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   357 			      Calc ("Groups.plus_class.plus", eval_binop "#add_")
   358 			      (*immediately repeated rewrite pushes
   359 					    '+' into precondition !*)
   360 			      ],
   361 		     scr = EmptyScr}, 
   362 	 srls = Erls, calc = [], errpatts = [],
   363 	 rules = [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   364 		  Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   365 		  Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
   366 		  Thm ("tl_Cons",num_str @{thm tl_Cons}),
   367 		  Thm ("tl_Nil",num_str @{thm tl_Nil}),
   368 		  Calc ("EqSystem.occur'_exactly'_in", 
   369 			eval_occur_exactly_in 
   370 			    "#eval_occur_exactly_in_")
   371 		  ],
   372 	 scr = EmptyScr};
   373 *}
   374 ML {*
   375 
   376 (*WN060914 quickly created for 4x4; 
   377  more similarity to prls_triangular desirable*)
   378 val prls_triangular4 = 
   379     Rls {id="prls_triangular4", preconds = [], 
   380 	 rew_ord = ("e_rew_ord", e_rew_ord), 
   381 	 erls = Rls {id="erls_prls_triangular4", preconds = [], 
   382 		     rew_ord = ("e_rew_ord", e_rew_ord), 
   383 		     erls = Erls, srls = Erls, calc = [], errpatts = [],
   384 		     rules = [(*for precond NTH_CONS ...*)
   385 			      Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   386 			      Calc ("Groups.plus_class.plus", eval_binop "#add_")
   387 			      (*immediately repeated rewrite pushes
   388 					    '+' into precondition !*)
   389 			      ],
   390 		     scr = EmptyScr}, 
   391 	 srls = Erls, calc = [], errpatts = [],
   392 	 rules = [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   393 		  Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   394 		  Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
   395 		  Thm ("tl_Cons",num_str @{thm tl_Cons}),
   396 		  Thm ("tl_Nil",num_str @{thm tl_Nil}),
   397 		  Calc ("EqSystem.occur'_exactly'_in", 
   398 			eval_occur_exactly_in 
   399 			    "#eval_occur_exactly_in_")
   400 		  ],
   401 	 scr = EmptyScr};
   402 *}
   403 
   404 ML {*
   405 ruleset' := 
   406 overwritelthy @{theory} 
   407 	      (!ruleset', 
   408 [("simplify_System_parenthesized", prep_rls simplify_System_parenthesized),
   409  ("simplify_System", prep_rls simplify_System),
   410  ("isolate_bdvs", prep_rls isolate_bdvs),
   411  ("isolate_bdvs_4x4", prep_rls isolate_bdvs_4x4),
   412  ("order_system", prep_rls order_system),
   413  ("order_add_mult_System", prep_rls order_add_mult_System),
   414  ("norm_System_noadd_fractions", prep_rls norm_System_noadd_fractions),
   415  ("norm_System", prep_rls norm_System)
   416  ]);
   417 *}
   418 
   419 ML {*
   420 (** problems **)
   421 
   422 store_pbt
   423  (prep_pbt thy "pbl_equsys" [] e_pblID
   424  (["system"],
   425   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   426    ("#Find"  ,["solution ss'''"](*''' is copy-named*))
   427   ],
   428   append_rls "e_rls" e_rls [(*for preds in where_*)], 
   429   SOME "solveSystem e_s v_s", 
   430   []));
   431 store_pbt
   432  (prep_pbt thy "pbl_equsys_lin" [] e_pblID
   433  (["linear", "system"],
   434   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   435    (*TODO.WN050929 check linearity*)
   436    ("#Find"  ,["solution ss'''"])
   437   ],
   438   append_rls "e_rls" e_rls [(*for preds in where_*)], 
   439   SOME "solveSystem e_s v_s", 
   440   []));
   441 store_pbt
   442  (prep_pbt thy "pbl_equsys_lin_2x2" [] e_pblID
   443  (["2x2", "linear", "system"],
   444   (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   445   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   446    ("#Where"  ,["LENGTH (e_s:: bool list) = 2", "LENGTH v_s = 2"]),
   447    ("#Find"  ,["solution ss'''"])
   448   ],
   449   append_rls "prls_2x2_linear_system" e_rls 
   450 			     [Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   451 			      Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   452 			      Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   453 			      Calc ("HOL.eq",eval_equal "#equal_")
   454 			      ], 
   455   SOME "solveSystem e_s v_s", 
   456   []));
   457 *}
   458 ML {*
   459 store_pbt
   460  (prep_pbt thy "pbl_equsys_lin_2x2_tri" [] e_pblID
   461  (["triangular", "2x2", "linear", "system"],
   462   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   463    ("#Where"  ,
   464     ["(tl v_s) from v_s occur_exactly_in (NTH 1 (e_s::bool list))",
   465      "    v_s  from v_s occur_exactly_in (NTH 2 (e_s::bool list))"]),
   466    ("#Find"  ,["solution ss'''"])
   467   ],
   468   prls_triangular, 
   469   SOME "solveSystem e_s v_s", 
   470   [["EqSystem","top_down_substitution","2x2"]]));
   471 store_pbt
   472  (prep_pbt thy "pbl_equsys_lin_2x2_norm" [] e_pblID
   473  (["normalize", "2x2", "linear", "system"],
   474   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   475    ("#Find"  ,["solution ss'''"])
   476   ],
   477   append_rls "e_rls" e_rls [(*for preds in where_*)], 
   478   SOME "solveSystem e_s v_s", 
   479   [["EqSystem","normalize","2x2"]]));
   480 store_pbt
   481  (prep_pbt thy "pbl_equsys_lin_3x3" [] e_pblID
   482  (["3x3", "linear", "system"],
   483   (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   484   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   485    ("#Where"  ,["LENGTH (e_s:: bool list) = 3", "LENGTH v_s = 3"]),
   486    ("#Find"  ,["solution ss'''"])
   487   ],
   488   append_rls "prls_3x3_linear_system" e_rls 
   489 			     [Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   490 			      Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   491 			      Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   492 			      Calc ("HOL.eq",eval_equal "#equal_")
   493 			      ], 
   494   SOME "solveSystem e_s v_s", 
   495   []));
   496 store_pbt
   497  (prep_pbt thy "pbl_equsys_lin_4x4" [] e_pblID
   498  (["4x4", "linear", "system"],
   499   (*~~~~~~~~~~~~~~~~~~~~~~~~~*)
   500   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   501    ("#Where"  ,["LENGTH (e_s:: bool list) = 4", "LENGTH v_s = 4"]),
   502    ("#Find"  ,["solution ss'''"])
   503   ],
   504   append_rls "prls_4x4_linear_system" e_rls 
   505 			     [Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   506 			      Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   507 			      Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   508 			      Calc ("HOL.eq",eval_equal "#equal_")
   509 			      ], 
   510   SOME "solveSystem e_s v_s", 
   511   []));
   512 *}
   513 ML {*
   514 store_pbt
   515  (prep_pbt thy "pbl_equsys_lin_4x4_tri" [] e_pblID
   516  (["triangular", "4x4", "linear", "system"],
   517   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   518    ("#Where" , (*accepts missing variables up to diagional form*)
   519     ["(NTH 1 (v_s::real list)) occurs_in (NTH 1 (e_s::bool list))",
   520      "(NTH 2 (v_s::real list)) occurs_in (NTH 2 (e_s::bool list))",
   521      "(NTH 3 (v_s::real list)) occurs_in (NTH 3 (e_s::bool list))",
   522      "(NTH 4 (v_s::real list)) occurs_in (NTH 4 (e_s::bool list))"
   523      ]),
   524    ("#Find"  ,["solution ss'''"])
   525   ],
   526   append_rls "prls_tri_4x4_lin_sys" prls_triangular
   527 	     [Calc ("Atools.occurs'_in",eval_occurs_in "")], 
   528   SOME "solveSystem e_s v_s", 
   529   [["EqSystem","top_down_substitution","4x4"]]));
   530 *}
   531 ML {*
   532 store_pbt
   533  (prep_pbt thy "pbl_equsys_lin_4x4_norm" [] e_pblID
   534  (["normalize", "4x4", "linear", "system"],
   535   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   536    (*LENGTH is checked 1 level above*)
   537    ("#Find"  ,["solution ss'''"])
   538   ],
   539   append_rls "e_rls" e_rls [(*for preds in where_*)], 
   540   SOME "solveSystem e_s v_s", 
   541   [["EqSystem","normalize","4x4"]]));
   542 
   543 
   544 (* show_ptyps();
   545    *)
   546 
   547 *}
   548 ML {*
   549 (** methods **)
   550 
   551 store_met
   552     (prep_met thy "met_eqsys" [] e_metID
   553 	      (["EqSystem"],
   554 	       [],
   555 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   556 		srls = Erls, prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   557 	       "empty_script"
   558 	       ));
   559 store_met
   560     (prep_met thy "met_eqsys_topdown" [] e_metID
   561 	      (["EqSystem","top_down_substitution"],
   562 	       [],
   563 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   564 		srls = Erls, prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   565 	       "empty_script"
   566 	       ));
   567 *}
   568 ML {*
   569 store_met
   570     (prep_met thy "met_eqsys_topdown_2x2" [] e_metID
   571 	 (["EqSystem","top_down_substitution","2x2"],
   572 	  [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   573 	   ("#Where"  ,
   574 	    ["(tl v_s) from v_s occur_exactly_in (NTH 1 (e_s::bool list))",
   575 	     "    v_s  from v_s occur_exactly_in (NTH 2 (e_s::bool list))"]),
   576 	   ("#Find"  ,["solution ss'''"])
   577 	   ],
   578 	  {rew_ord'="ord_simplify_System", rls' = Erls, calc = [], 
   579 	   srls = append_rls "srls_top_down_2x2" e_rls
   580 				  [Thm ("hd_thm",num_str @{thm hd_thm}),
   581 				   Thm ("tl_Cons",num_str @{thm tl_Cons}),
   582 				   Thm ("tl_Nil",num_str @{thm tl_Nil})
   583 				   ], 
   584 	   prls = prls_triangular, crls = Erls, errpats = [], nrls = Erls},
   585 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   586 "  (let e_1 = Take (hd e_s);                                                " ^
   587 "       e_1 = ((Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   588 "                                  isolate_bdvs False))     @@               " ^
   589 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   590 "                                  simplify_System False))) e_1;            " ^
   591 "       e_2 = Take (hd (tl e_s));                                           " ^
   592 "       e_2 = ((Substitute [e_1]) @@                                       " ^
   593 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   594 "                                  simplify_System_parenthesized False)) @@  " ^
   595 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   596 "                                  isolate_bdvs False))     @@               " ^
   597 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   598 "                                  simplify_System False))) e_2;            " ^
   599 "       e__s = Take [e_1, e_2]                                             " ^
   600 "   in (Try (Rewrite_Set order_system False)) e__s)"
   601 (*---------------------------------------------------------------------------
   602   this script does NOT separate the equations as above, 
   603   but it does not yet work due to preliminary script-interpreter,
   604   see eqsystem.sml 'script [EqSystem,top_down_substitution,2x2] Vers.2'
   605 
   606 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =         " ^
   607 "  (let e__s = Take e_s;                                              " ^
   608 "       e_1 = hd e__s;                                               " ^
   609 "       e_2 = hd (tl e__s);                                          " ^
   610 "       e__s = [e_1, Substitute [e_1] e_2]                         " ^
   611 "   in ((Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   612 "                                  simplify_System_parenthesized False)) @@ " ^
   613 "       (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))] " ^
   614 "                              isolate_bdvs False))              @@   " ^
   615 "       (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   616 "                                  simplify_System False))) e__s)"
   617 ---------------------------------------------------------------------------*)
   618 	  ));
   619 *}
   620 ML {*
   621 store_met
   622     (prep_met thy "met_eqsys_norm" [] e_metID
   623 	      (["EqSystem","normalize"],
   624 	       [],
   625 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   626 		srls = Erls, prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   627 	       "empty_script"
   628 	       ));
   629 *}
   630 ML {*
   631 store_met
   632     (prep_met thy "met_eqsys_norm_2x2" [] e_metID
   633 	      (["EqSystem","normalize","2x2"],
   634 	       [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   635 		("#Find"  ,["solution ss'''"])],
   636 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   637 		srls = append_rls "srls_normalize_2x2" e_rls
   638 				  [Thm ("hd_thm",num_str @{thm hd_thm}),
   639 				   Thm ("tl_Cons",num_str @{thm tl_Cons}),
   640 				   Thm ("tl_Nil",num_str @{thm tl_Nil})
   641 				   ], 
   642 		prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   643 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   644 "  (let e__s = ((Try (Rewrite_Set norm_Rational False)) @@                   " ^
   645 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   646 "                                  simplify_System_parenthesized False)) @@  " ^
   647 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   648 "                                                    isolate_bdvs False)) @@ " ^
   649 "               (Try (Rewrite_Set_Inst [(bdv_1, hd v_s),(bdv_2, hd (tl v_s))]" ^
   650 "                                  simplify_System_parenthesized False)) @@  " ^
   651 "               (Try (Rewrite_Set order_system False))) e_s                  " ^
   652 "   in (SubProblem (EqSystem',[linear,system],[no_met])                      " ^
   653 "                  [BOOL_LIST e__s, REAL_LIST v_s]))"
   654 	       ));
   655 
   656 *}
   657 ML {*
   658 (*this is for NTH only*)
   659 val srls = Rls {id="srls_normalize_4x4", 
   660 		preconds = [], 
   661 		rew_ord = ("termlessI",termlessI), 
   662 		erls = append_rls "erls_in_srls_IntegrierenUnd.." e_rls
   663 				  [(*for asm in NTH_CONS ...*)
   664 				   Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   665 				   (*2nd NTH_CONS pushes n+-1 into asms*)
   666 				   Calc("Groups.plus_class.plus", eval_binop "#add_")
   667 				   ], 
   668 		srls = Erls, calc = [], errpatts = [],
   669 		rules = [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   670 			 Calc("Groups.plus_class.plus", eval_binop "#add_"),
   671 			 Thm ("NTH_NIL",num_str @{thm NTH_NIL})],
   672 		scr = EmptyScr};
   673 
   674 store_met
   675     (prep_met thy "met_eqsys_norm_4x4" [] e_metID
   676 	      (["EqSystem","normalize","4x4"],
   677 	       [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   678 		("#Find"  ,["solution ss'''"])],
   679 	       {rew_ord'="tless_true", rls' = Erls, calc = [], 
   680 		srls = append_rls "srls_normalize_4x4" srls
   681 				  [Thm ("hd_thm",num_str @{thm hd_thm}),
   682 				   Thm ("tl_Cons",num_str @{thm tl_Cons}),
   683 				   Thm ("tl_Nil",num_str @{thm tl_Nil})
   684 				   ], 
   685 		prls = Erls, crls = Erls, errpats = [], nrls = Erls},
   686 (*STOPPED.WN06? met ["EqSystem","normalize","4x4"] @@@@@@@@@@@@@@@@@@@@@@@@@@@*)
   687 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   688 "  (let e__s =                                                               " ^
   689 "     ((Try (Rewrite_Set norm_Rational False)) @@                            " ^
   690 "      (Repeat (Rewrite commute_0_equality False)) @@                        " ^
   691 "      (Try (Rewrite_Set_Inst [(bdv_1, NTH 1 v_s),(bdv_2, NTH 2 v_s ),     " ^
   692 "                              (bdv_3, NTH 3 v_s),(bdv_3, NTH 4 v_s )]     " ^
   693 "                             simplify_System_parenthesized False))    @@    " ^
   694 "      (Try (Rewrite_Set_Inst [(bdv_1, NTH 1 v_s),(bdv_2, NTH 2 v_s ),     " ^
   695 "                              (bdv_3, NTH 3 v_s),(bdv_3, NTH 4 v_s )]     " ^
   696 "                             isolate_bdvs_4x4 False))                 @@    " ^
   697 "      (Try (Rewrite_Set_Inst [(bdv_1, NTH 1 v_s),(bdv_2, NTH 2 v_s ),     " ^
   698 "                              (bdv_3, NTH 3 v_s),(bdv_3, NTH 4 v_s )]     " ^
   699 "                             simplify_System_parenthesized False))    @@    " ^
   700 "      (Try (Rewrite_Set order_system False)))                           e_s " ^
   701 "   in (SubProblem (EqSystem',[linear,system],[no_met])                      " ^
   702 "                  [BOOL_LIST e__s, REAL_LIST v_s]))"
   703 ));
   704 
   705 store_met
   706 (prep_met thy "met_eqsys_topdown_4x4" [] e_metID
   707 	  (["EqSystem","top_down_substitution","4x4"],
   708 	   [("#Given" ,["equalities e_s", "solveForVars v_s"]),
   709 	    ("#Where" , (*accepts missing variables up to diagonal form*)
   710 	     ["(NTH 1 (v_s::real list)) occurs_in (NTH 1 (e_s::bool list))",
   711 	      "(NTH 2 (v_s::real list)) occurs_in (NTH 2 (e_s::bool list))",
   712 	      "(NTH 3 (v_s::real list)) occurs_in (NTH 3 (e_s::bool list))",
   713 	      "(NTH 4 (v_s::real list)) occurs_in (NTH 4 (e_s::bool list))"
   714 	      ]),
   715 	    ("#Find"  ,["solution ss'''"])
   716 	    ],
   717 	   {rew_ord'="ord_simplify_System", rls' = Erls, calc = [], 
   718 	    srls = append_rls "srls_top_down_4x4" srls [], 
   719 	    prls = append_rls "prls_tri_4x4_lin_sys" prls_triangular
   720 			      [Calc ("Atools.occurs'_in",eval_occurs_in "")], 
   721 	    crls = Erls, errpats = [], nrls = Erls},
   722 (*FIXXXXME.WN060916: this script works ONLY for exp 7.79 @@@@@@@@@@@@@@@@@@@@*)
   723 "Script SolveSystemScript (e_s::bool list) (v_s::real list) =                " ^
   724 "  (let e_1 = NTH 1 e_s;                                                    " ^
   725 "       e_2 = Take (NTH 2 e_s);                                             " ^
   726 "       e_2 = ((Substitute [e_1]) @@                                         " ^
   727 "              (Try (Rewrite_Set_Inst [(bdv_1,NTH 1 v_s),(bdv_2,NTH 2 v_s)," ^
   728 "                                      (bdv_3,NTH 3 v_s),(bdv_4,NTH 4 v_s)]" ^
   729 "                                 simplify_System_parenthesized False)) @@   " ^
   730 "              (Try (Rewrite_Set_Inst [(bdv_1,NTH 1 v_s),(bdv_2,NTH 2 v_s)," ^
   731 "                                      (bdv_3,NTH 3 v_s),(bdv_4,NTH 4 v_s)]" ^
   732 "                                 isolate_bdvs False))                  @@   " ^
   733 "              (Try (Rewrite_Set_Inst [(bdv_1,NTH 1 v_s),(bdv_2,NTH 2 v_s)," ^
   734 "                                      (bdv_3,NTH 3 v_s),(bdv_4,NTH 4 v_s)]" ^
   735 "                                 norm_Rational False)))             e_2     " ^
   736 "    in [e_1, e_2, NTH 3 e_s, NTH 4 e_s])"
   737 ));
   738 *}
   739 
   740 end