src/Tools/isac/Knowledge/EqSystem.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 04 Aug 2022 12:48:37 +0200
changeset 60509 2e0b7ca391dc
parent 60506 145e45cd7a0f
child 60515 03e19793d81e
permissions -rw-r--r--
polish naming in Rewrite_Order
     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 lemma commute_0_equality : \<open>(0 = a) = (a = 0)\<close>
    22   by auto
    23 
    24 axiomatization where
    25   (*WN0510 see simliar rules 'isolate_' 'separate_' (by RL)
    26     [bdv_1,bdv_2,bdv_3,bdv_4] work also for 2 and 3 bdvs, ugly !*)
    27   (* these require lemmas about occur_exactly_in* )
    28   lemma xxx : 
    29     \<open>[| [] from [bdv_1, bdv_2, bdv_3, bdv_4] occur_exactly_in a |] ==> (a + b = c) = (b = c + -1 * a)\<close>
    30     by auto
    31   ( **)
    32   separate_bdvs_add:   
    33     "[| [] from [bdv_1,bdv_2,bdv_3,bdv_4] occur_exactly_in a |] 
    34 		      			     ==> (a + b = c) = (b = c + -1*a)" and
    35   separate_bdvs0:
    36     "[| some_of [bdv_1,bdv_2,bdv_3,bdv_4] occur_in b; Not (b=!=0)  |] 
    37 		      			     ==> (a = b) = (a + -1*b = 0)" and
    38   separate_bdvs_add1:  
    39     "[| some_of [bdv_1,bdv_2,bdv_3,bdv_4] occur_in c |] 
    40 		      			     ==> (a = b + c) = (a + -1*c = b)" and
    41   separate_bdvs_add2:
    42     "[| Not (some_of [bdv_1,bdv_2,bdv_3,bdv_4] occur_in a) |] 
    43 		      			     ==> (a + b = c) = (b = -1*a + c)" and
    44   separate_bdvs_mult:  
    45     "[| [] from [bdv_1,bdv_2,bdv_3,bdv_4] occur_exactly_in a; Not (a=!=0) |] 
    46 		      			     ==>(a * b = c) = (b = c / a)" 
    47 axiomatization where (*..if replaced by "and" we get an error in 
    48   ---  rewrite in [EqSystem,normalise,2x2] --- step "--- 3---";*)
    49   order_system_NxN:     "[a,b] = [b,a]"
    50   (*requires rew_ord for termination, eg. ord_simplify_Integral;
    51     works for lists of any length, interestingly !?!
    52     CONTRADICTS PROPERTIES OF LIST: take set instead*)
    53 
    54 ML \<open>
    55 (** eval functions **)
    56 
    57 (*certain variables of a given list occur _all_ in a term
    58   args: all: ..variables, which are under consideration (eg. the bound vars)
    59         vs:  variables which must be in t, 
    60              and none of the others in all must be in t
    61         t: the term under consideration
    62  *)
    63 fun occur_exactly_in vs all t =
    64     let fun occurs_in' a b = Prog_Expr.occurs_in b a
    65     in foldl and_ (true, map (occurs_in' t) vs)
    66        andalso not (foldl or_ (false, map (occurs_in' t) 
    67                                           (subtract op = vs all)))
    68     end;
    69 
    70 (*("occur_exactly_in", ("EqSystem.occur_exactly_in", 
    71                         eval_occur_exactly_in "#eval_occur_exactly_in_") )*)
    72 fun eval_occur_exactly_in _ "EqSystem.occur_exactly_in"
    73 			  (p as (Const (\<^const_name>\<open>EqSystem.occur_exactly_in\<close>,_) 
    74 				       $ vs $ all $ t)) _ =
    75     if occur_exactly_in (TermC.isalist2list vs) (TermC.isalist2list all) t
    76     then SOME ((UnparseC.term p) ^ " = True",
    77 	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
    78     else SOME ((UnparseC.term p) ^ " = False",
    79 	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
    80   | eval_occur_exactly_in _ _ _ _ = NONE;
    81 \<close>
    82 calculation occur_exactly_in = \<open>eval_occur_exactly_in "#eval_occur_exactly_in_"\<close>
    83 
    84 ML \<open>
    85 (** rewrite order 'ord_simplify_System' **)
    86 
    87 (* order wrt. several linear (i.e. without exponents) variables "c", "c_2",..
    88    which leaves the monomials containing c, c_2,... at the end of an Integral
    89    and puts the c, c_2,... rightmost within a monomial.
    90 
    91    WN050906 this is a quick and dirty adaption of ord_make_polynomial_in,
    92    which was most adequate, because it uses size_of_term*)
    93 (**)
    94 local (*. for simplify_System .*)
    95 (**)
    96 open Term;  (* for type order = EQUAL | LESS | GREATER *)
    97 
    98 fun pr_ord EQUAL = "EQUAL"
    99   | pr_ord LESS  = "LESS"
   100   | pr_ord GREATER = "GREATER";
   101 
   102 fun dest_hd' (Const (a, T)) = (((a, 0), T), 0)
   103   | dest_hd' (Free (ccc, T)) =
   104     (case Symbol.explode ccc of
   105 	"c"::[] => ((("|||||||||||||||||||||", 0), T), 1)(*greatest string WN*)
   106       | "c"::"_"::_ => ((("|||||||||||||||||||||", 0), T), 1)
   107       | _ => (((ccc, 0), T), 1))
   108   | dest_hd' (Var v) = (v, 2)
   109   | dest_hd' (Bound i) = ((("", i), dummyT), 3)
   110   | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4)
   111   | dest_hd' _ = raise ERROR "dest_hd': uncovered case in fun.def.";
   112 
   113 fun size_of_term' (Free (ccc, _)) =
   114     (case Symbol.explode ccc of (*WN0510 hack for the bound variables*)
   115 	"c"::[] => 1000
   116       | "c"::"_"::is => 1000 * ((TermC.int_of_str o implode) is)
   117       | _ => 1)
   118   | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
   119   | size_of_term' (f$t) = size_of_term' f  +  size_of_term' t
   120   | size_of_term' _ = 1;
   121 
   122 fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
   123     (case term_ord' pr thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord)
   124   | term_ord' pr thy (t, u) =
   125     (if pr
   126      then 
   127        let
   128          val (f, ts) = strip_comb t and (g, us) = strip_comb u;
   129          val _ = tracing ("t= f @ ts= \"" ^ UnparseC.term_in_thy thy f ^ "\" @ \"[" ^
   130            commas (map (UnparseC.term_in_thy thy) ts) ^ "]\"");
   131          val _ = tracing ("u= g @ us= \"" ^ UnparseC.term_in_thy thy g ^ "\" @ \"[" ^
   132            commas (map (UnparseC.term_in_thy thy) us) ^ "]\"");
   133          val _ = tracing ("size_of_term (t, u) = (" ^ string_of_int (size_of_term' t) ^ ", " ^
   134            string_of_int (size_of_term' u) ^ ")");
   135          val _ = tracing ("hd_ord (f, g)      = " ^ ((pr_ord o hd_ord) (f, g)) );
   136     (** )val _ = @{print tracing}{a = "hd_ord (f, g)      = ", z = hd_ord (f, g)}( **)
   137          val _ = tracing ("terms_ord (ts, us) = " ^(pr_ord o terms_ord str false) (ts,us));
   138          val _= tracing ("-------");
   139        in () end
   140      else ();
   141     case int_ord (size_of_term' t, size_of_term' u) of
   142       EQUAL =>
   143         let val (f, ts) = strip_comb t and (g, us) = strip_comb u 
   144         in (case hd_ord (f, g) of 
   145               EQUAL => (terms_ord str pr) (ts, us) 
   146             | ord => ord)
   147         end
   148 	 | ord => ord)
   149 and hd_ord (f, g) =                                        (* ~ term.ML *)
   150   prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord (dest_hd' f, dest_hd' g)
   151 and terms_ord _ pr (ts, us) = list_ord (term_ord' pr (ThyC.get_theory "Isac_Knowledge"))(ts, us);
   152 (**)
   153 in
   154 (**)
   155 (*WN0510 for preliminary use in eval_order_system, see case-study mat-eng.tex
   156 fun ord_simplify_System_rev (pr:bool) thy subst tu = 
   157     (term_ord' pr thy (Library.swap tu) = LESS);*)
   158 
   159 (*for the rls's*)
   160 fun ord_simplify_System (pr:bool) thy _(*subst*) (ts, us) = 
   161     (term_ord' pr thy (TermC.numerals_to_Free ts, TermC.numerals_to_Free us) = LESS);
   162 
   163 (**)end;(**)
   164 \<close> ML \<open>
   165 \<close>
   166 setup \<open>KEStore_Elems.add_rew_ord [
   167   ("ord_simplify_System", ord_simplify_System false \<^theory>)]\<close>
   168 
   169 ML \<open>
   170 (** rulesets **)
   171 
   172 (*.adapted from 'order_add_mult_in' by just replacing the rew_ord.*)
   173 val order_add_mult_System = 
   174   Rule_Def.Repeat{
   175     id = "order_add_mult_System", preconds = [], 
   176     rew_ord = ("ord_simplify_System", ord_simplify_System false @{theory "Integrate"}),
   177     erls = Rule_Set.empty,srls = Rule_Set.Empty, calc = [], errpatts = [],
   178     rules = [
   179       \<^rule_thm>\<open>mult.commute\<close>, (* z * w = w * z *)
   180       \<^rule_thm>\<open>real_mult_left_commute\<close>, (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
   181       \<^rule_thm>\<open>mult.assoc\<close>,	 (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
   182       \<^rule_thm>\<open>add.commute\<close>, (*z + w = w + z*)
   183       \<^rule_thm>\<open>add.left_commute\<close>, (*x + (y + z) = y + (x + z)*)
   184       \<^rule_thm>\<open>add.assoc\<close>	],  (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
   185     scr = Rule.Empty_Prog};
   186 \<close>
   187 ML \<open>
   188 (*.adapted from 'norm_Rational' by
   189   #1 using 'ord_simplify_System' in 'order_add_mult_System'
   190   #2 NOT using common_nominator_p                          .*)
   191 val norm_System_noadd_fractions = 
   192   Rule_Def.Repeat {id = "norm_System_noadd_fractions", preconds = [], 
   193     rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), 
   194     erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   195     rules = [(*sequence given by operator precedence*)
   196   		Rule.Rls_ discard_minus,
   197   		Rule.Rls_ powers,
   198   		Rule.Rls_ rat_mult_divide,
   199   		Rule.Rls_ expand,
   200   		Rule.Rls_ reduce_0_1_2,
   201   		Rule.Rls_ (*order_add_mult #1*) order_add_mult_System,
   202   		Rule.Rls_ collect_numerals,
   203   		(*Rule.Rls_ add_fractions_p, #2*)
   204   		Rule.Rls_ cancel_p],
   205     scr = Rule.Empty_Prog};
   206 \<close>
   207 ML \<open>
   208 (*.adapted from 'norm_Rational' by
   209   *1* using 'ord_simplify_System' in 'order_add_mult_System'.*)
   210 val norm_System = 
   211   Rule_Def.Repeat {id = "norm_System", preconds = [], 
   212     rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), 
   213     erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   214     rules = [(*sequence given by operator precedence*)
   215   		Rule.Rls_ discard_minus,
   216   		Rule.Rls_ powers,
   217   		Rule.Rls_ rat_mult_divide,
   218   		Rule.Rls_ expand,
   219   		Rule.Rls_ reduce_0_1_2,
   220   		Rule.Rls_ (*order_add_mult *1*) order_add_mult_System,
   221   		Rule.Rls_ collect_numerals,
   222   		Rule.Rls_ add_fractions_p,
   223   		Rule.Rls_ cancel_p],
   224     scr = Rule.Empty_Prog};
   225 \<close>
   226 ML \<open>
   227 (*.simplify an equational system BEFORE solving it such that parentheses are
   228    ( ((u0*v0)*w0) + ( ((u1*v1)*w1) * c + ... +((u4*v4)*w4) * c_4 ) )
   229 ATTENTION: works ONLY for bound variables c, c_1, c_2, c_3, c_4 :ATTENTION
   230    This is a copy from 'make_ratpoly_in' with respective reductions:
   231    *0* expand the term, ie. distribute * and / over +
   232    *1* ord_simplify_System instead of termlessI
   233    *2* no add_fractions_p (= common_nominator_p_rls !)
   234    *3* discard_parentheses only for (.*(.*.))
   235    analoguous to simplify_Integral                                       .*)
   236 val simplify_System_parenthesized = 
   237   Rule_Set.Sequence {id = "simplify_System_parenthesized", preconds = []:term list, 
   238     rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
   239     erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   240     rules = [
   241        \<^rule_thm>\<open>distrib_right\<close>, (*"(?z1.0 + ?z2.0) * ?w = ?z1.0 * ?w + ?z2.0 * ?w"*)
   242 	     \<^rule_thm>\<open>add_divide_distrib\<close>, (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)
   243 	     Rule.Rls_ norm_Rational_noadd_fractions,
   244 	     Rule.Rls_ (*order_add_mult_in*) norm_System_noadd_fractions,
   245 	     \<^rule_thm_sym>\<open>mult.assoc\<close>,
   246 	     Rule.Rls_ collect_bdv, (*from make_polynomial_in WN051031 welldone?*)
   247 	     Rule.Rls_ separate_bdv2,
   248 	     \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e")],
   249     scr = Rule.Empty_Prog};      
   250 \<close>
   251 ML \<open>
   252 (*.simplify an equational system AFTER solving it;
   253    This is a copy of 'make_ratpoly_in' with the differences
   254    *1* ord_simplify_System instead of termlessI           .*)
   255 (*TODO.WN051031 ^^^^^^^^^^ should be in EACH rls contained *)
   256 val simplify_System = 
   257   Rule_Set.Sequence {id = "simplify_System", preconds = []:term list, 
   258     rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
   259     erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   260     rules = [
   261       Rule.Rls_ norm_Rational,
   262 	    Rule.Rls_ (*order_add_mult_in*) norm_System (**1**),
   263 	    Rule.Rls_ discard_parentheses,
   264 	    Rule.Rls_ collect_bdv, (*from make_polynomial_in WN051031 welldone?*)
   265 	    Rule.Rls_ separate_bdv2,
   266 	    \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e")],
   267     scr = Rule.Empty_Prog};      
   268 (*
   269 val simplify_System = 
   270     Rule_Set.append_rules "simplify_System" simplify_System_parenthesized
   271 	       [\<^rule_thm_sym>\<open>add.assoc\<close>];
   272 *)
   273 \<close>
   274 ML \<open>
   275 val isolate_bdvs = 
   276   Rule_Def.Repeat {
   277     id="isolate_bdvs", preconds = [], rew_ord = ("Rewrite_Ord.id_empty", Rewrite_Ord.function_empty), 
   278     erls = Rule_Set.append_rules "erls_isolate_bdvs" Rule_Set.empty [
   279       (\<^rule_eval>\<open>occur_exactly_in\<close> (eval_occur_exactly_in "#eval_occur_exactly_in_"))], 
   280     srls = Rule_Set.Empty, calc = [], errpatts = [],
   281     rules = [
   282       \<^rule_thm>\<open>commute_0_equality\<close>,
   283       \<^rule_thm>\<open>separate_bdvs_add\<close>,
   284       \<^rule_thm>\<open>separate_bdvs_mult\<close>],
   285     scr = Rule.Empty_Prog};
   286 \<close>
   287 ML \<open>
   288 val isolate_bdvs_4x4 = 
   289   Rule_Def.Repeat {id="isolate_bdvs_4x4", preconds = [], 
   290     rew_ord = ("Rewrite_Ord.id_empty", Rewrite_Ord.function_empty), 
   291     erls = Rule_Set.append_rules "erls_isolate_bdvs_4x4" Rule_Set.empty [
   292       \<^rule_eval>\<open>occur_exactly_in\<close> (eval_occur_exactly_in "#eval_occur_exactly_in_"),
   293       \<^rule_eval>\<open>Prog_Expr.ident\<close> (Prog_Expr.eval_ident "#ident_"),
   294       \<^rule_eval>\<open>Prog_Expr.some_occur_in\<close> (Prog_Expr.eval_some_occur_in "#some_occur_in_"),
   295       \<^rule_thm>\<open>not_true\<close>,
   296       \<^rule_thm>\<open>not_false\<close>], 
   297     srls = Rule_Set.Empty, calc = [], errpatts = [],
   298     rules = [
   299       \<^rule_thm>\<open>commute_0_equality\<close>,
   300       \<^rule_thm>\<open>separate_bdvs0\<close>,
   301       \<^rule_thm>\<open>separate_bdvs_add1\<close>,
   302       \<^rule_thm>\<open>separate_bdvs_add2\<close>,
   303       \<^rule_thm>\<open>separate_bdvs_mult\<close>],
   304     scr = Rule.Empty_Prog};
   305 
   306 \<close>
   307 ML \<open>
   308 
   309 (*.order the equations in a system such, that a triangular system (if any)
   310    appears as [..c_4 = .., ..., ..., ..c_1 + ..c_2 + ..c_3 ..c_4 = ..].*)
   311 val order_system = 
   312   Rule_Def.Repeat {id="order_system", preconds = [], 
   313 	  rew_ord = ("ord_simplify_System", ord_simplify_System false \<^theory>), 
   314 	  erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   315 	  rules = [
   316       \<^rule_thm>\<open>order_system_NxN\<close>],
   317 	  scr = Rule.Empty_Prog};
   318 
   319 val prls_triangular = 
   320   Rule_Def.Repeat {
   321     id="prls_triangular", preconds = [], rew_ord = ("Rewrite_Ord.id_empty", Rewrite_Ord.function_empty), 
   322     erls = Rule_Def.Repeat {
   323       id="erls_prls_triangular", preconds = [], rew_ord = ("Rewrite_Ord.id_empty", Rewrite_Ord.function_empty), 
   324       erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   325       rules = [(*for precond NTH_CONS ...*)
   326          \<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_"),
   327          \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
   328          \<^rule_eval>\<open>occur_exactly_in\<close> (eval_occur_exactly_in "#eval_occur_exactly_in_")],
   329          (*immediately repeated rewrite pushes '+' into precondition !*)
   330       scr = Rule.Empty_Prog}, 
   331     srls = Rule_Set.Empty, calc = [], errpatts = [],
   332     rules = [
   333       \<^rule_thm>\<open>NTH_CONS\<close>,
   334       \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
   335       \<^rule_thm>\<open>NTH_NIL\<close>,
   336       \<^rule_thm>\<open>tl_Cons\<close>,
   337       \<^rule_thm>\<open>tl_Nil\<close>,
   338       \<^rule_eval>\<open>occur_exactly_in\<close> (eval_occur_exactly_in "#eval_occur_exactly_in_")],
   339     scr = Rule.Empty_Prog};
   340 \<close>
   341 ML \<open>
   342 
   343 (*WN060914 quickly created for 4x4; 
   344  more similarity to prls_triangular desirable*)
   345 val prls_triangular4 = 
   346   Rule_Def.Repeat {
   347   id="prls_triangular4", preconds = [], rew_ord = ("Rewrite_Ord.id_empty", Rewrite_Ord.function_empty), 
   348   erls = Rule_Def.Repeat {
   349     id="erls_prls_triangular4", preconds = [], rew_ord = ("Rewrite_Ord.id_empty", Rewrite_Ord.function_empty), 
   350     erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   351     rules = [(*for precond NTH_CONS ...*)
   352   	   \<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_"),
   353   	   \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_")],
   354   	   (*immediately repeated rewrite pushes '+' into precondition !*)
   355     scr = Rule.Empty_Prog}, 
   356   srls = Rule_Set.Empty, calc = [], errpatts = [],
   357   rules = [
   358     \<^rule_thm>\<open>NTH_CONS\<close>,
   359     \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
   360     \<^rule_thm>\<open>NTH_NIL\<close>,
   361     \<^rule_thm>\<open>tl_Cons\<close>,
   362     \<^rule_thm>\<open>tl_Nil\<close>,
   363     \<^rule_eval>\<open>occur_exactly_in\<close> (eval_occur_exactly_in "#eval_occur_exactly_in_")],
   364   scr = Rule.Empty_Prog};
   365 \<close>
   366 
   367 rule_set_knowledge
   368   simplify_System_parenthesized = \<open>prep_rls' simplify_System_parenthesized\<close> and
   369   simplify_System = \<open>prep_rls' simplify_System\<close> and
   370   isolate_bdvs = \<open>prep_rls' isolate_bdvs\<close> and
   371   isolate_bdvs_4x4 = \<open>prep_rls' isolate_bdvs_4x4\<close> and 
   372   order_system = \<open>prep_rls' order_system\<close> and 
   373   order_add_mult_System = \<open>prep_rls' order_add_mult_System\<close> and
   374   norm_System_noadd_fractions = \<open>prep_rls' norm_System_noadd_fractions\<close> and
   375   norm_System = \<open>prep_rls' norm_System\<close>
   376 
   377 
   378 section \<open>Problems\<close>
   379 
   380 problem pbl_equsys : "system" =
   381   \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
   382   CAS: "solveSystem e_s v_s"
   383   Given: "equalities e_s" "solveForVars v_s"
   384   Find: "solution ss'''" (*''' is copy-named*)
   385 
   386 problem pbl_equsys_lin : "LINEAR/system" =
   387   \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
   388   CAS: "solveSystem e_s v_s"
   389   Given: "equalities e_s" "solveForVars v_s"
   390   (*TODO.WN050929 check linearity*)
   391   Find: "solution ss'''"
   392 
   393 problem pbl_equsys_lin_2x2: "2x2/LINEAR/system" =
   394   \<open>Rule_Set.append_rules "prls_2x2_linear_system" Rule_Set.empty 
   395     [\<^rule_thm>\<open>LENGTH_CONS\<close>,
   396       \<^rule_thm>\<open>LENGTH_NIL\<close>,
   397       \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
   398       \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_")]\<close>
   399   CAS: "solveSystem e_s v_s"
   400   Given: "equalities e_s" "solveForVars v_s"
   401   Where: "Length (e_s:: bool list) = 2" "Length v_s = 2"
   402   Find: "solution ss'''"
   403 
   404 problem pbl_equsys_lin_2x2_tri : "triangular/2x2/LINEAR/system" =
   405   \<open>prls_triangular\<close>
   406   Method_Ref: "EqSystem/top_down_substitution/2x2"
   407   CAS: "solveSystem e_s v_s"
   408   Given: "equalities e_s" "solveForVars v_s"
   409   Where:
   410     "(tl v_s) from v_s occur_exactly_in (NTH 1 (e_s::bool list))"
   411     "    v_s  from v_s occur_exactly_in (NTH 2 (e_s::bool list))"
   412   Find: "solution ss'''"
   413 
   414 problem pbl_equsys_lin_2x2_norm : "normalise/2x2/LINEAR/system" =
   415   \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
   416   Method_Ref: "EqSystem/normalise/2x2"
   417   CAS: "solveSystem e_s v_s"
   418   Given: "equalities e_s" "solveForVars v_s"
   419   Find: "solution ss'''"
   420 
   421 problem pbl_equsys_lin_3x3 : "3x3/LINEAR/system" =
   422   \<open>Rule_Set.append_rules "prls_3x3_linear_system" Rule_Set.empty 
   423     [\<^rule_thm>\<open>LENGTH_CONS\<close>,
   424       \<^rule_thm>\<open>LENGTH_NIL\<close>,
   425       \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
   426       \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_")]\<close>
   427   CAS: "solveSystem e_s v_s"
   428   Given: "equalities e_s" "solveForVars v_s"
   429   Where: "Length (e_s:: bool list) = 3" "Length v_s = 3"
   430   Find: "solution ss'''"
   431 
   432 problem pbl_equsys_lin_4x4 : "4x4/LINEAR/system" =
   433   \<open>Rule_Set.append_rules "prls_4x4_linear_system" Rule_Set.empty 
   434     [\<^rule_thm>\<open>LENGTH_CONS\<close>,
   435       \<^rule_thm>\<open>LENGTH_NIL\<close>,
   436       \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
   437       \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_")]\<close>
   438   CAS: "solveSystem e_s v_s"
   439   Given: "equalities e_s" "solveForVars v_s"
   440   Where: "Length (e_s:: bool list) = 4" "Length v_s = 4"
   441   Find: "solution ss'''"
   442 
   443 problem pbl_equsys_lin_4x4_tri : "triangular/4x4/LINEAR/system" =
   444   \<open>Rule_Set.append_rules "prls_tri_4x4_lin_sys" prls_triangular
   445     [\<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in "")]\<close>
   446   Method_Ref: "EqSystem/top_down_substitution/4x4"
   447   CAS: "solveSystem e_s v_s"
   448   Given: "equalities e_s" "solveForVars v_s"
   449   Where: (*accepts missing variables up to diagional form*)
   450     "(NTH 1 (v_s::real list)) occurs_in (NTH 1 (e_s::bool list))"
   451     "(NTH 2 (v_s::real list)) occurs_in (NTH 2 (e_s::bool list))"
   452     "(NTH 3 (v_s::real list)) occurs_in (NTH 3 (e_s::bool list))"
   453     "(NTH 4 (v_s::real list)) occurs_in (NTH 4 (e_s::bool list))"
   454   Find: "solution ss'''"
   455 
   456 problem pbl_equsys_lin_4x4_norm : "normalise/4x4/LINEAR/system" =
   457   \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
   458   Method_Ref: "EqSystem/normalise/4x4"
   459   CAS: "solveSystem e_s v_s"
   460   Given: "equalities e_s" "solveForVars v_s"
   461   (*Length is checked 1 level above*)
   462   Find: "solution ss'''"
   463 
   464 ML \<open>
   465 (*this is for NTH only*)
   466 val srls = Rule_Def.Repeat {id="srls_normalise_4x4", 
   467 		preconds = [], 
   468 		rew_ord = ("termlessI",termlessI), 
   469 		erls = Rule_Set.append_rules "erls_in_srls_IntegrierenUnd.." Rule_Set.empty
   470 				  [(*for asm in NTH_CONS ...*)
   471 				   \<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_"),
   472 				   (*2nd NTH_CONS pushes n+-1 into asms*)
   473 				   \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_")
   474 				   ], 
   475 		srls = Rule_Set.Empty, calc = [], errpatts = [],
   476 		rules = [\<^rule_thm>\<open>NTH_CONS\<close>,
   477 			 \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
   478 			 \<^rule_thm>\<open>NTH_NIL\<close>],
   479 		scr = Rule.Empty_Prog};
   480 \<close>
   481 
   482 section \<open>Methods\<close>
   483 
   484 method met_eqsys : "EqSystem" =
   485   \<open>{rew_ord'="tless_true", rls' = Rule_Set.Empty, calc = [], srls = Rule_Set.Empty, prls = Rule_Set.Empty, crls = Rule_Set.Empty,
   486     errpats = [], nrls = Rule_Set.Empty}\<close>
   487 
   488 method met_eqsys_topdown : "EqSystem/top_down_substitution" =
   489   \<open>{rew_ord'="tless_true", rls' = Rule_Set.Empty, calc = [], srls = Rule_Set.Empty, prls = Rule_Set.Empty, crls = Rule_Set.Empty,
   490     errpats = [], nrls = Rule_Set.Empty}\<close>
   491 
   492 partial_function (tailrec) solve_system :: "bool list => real list => bool list"
   493   where
   494 "solve_system e_s v_s = (
   495   let
   496     e_1 = Take (hd e_s);                                                         
   497     e_1 = (
   498       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))] ''isolate_bdvs'')) #>                   
   499       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))] ''simplify_System''))
   500       ) e_1;                 
   501     e_2 = Take (hd (tl e_s));                                                    
   502     e_2 = (
   503       (Substitute [e_1]) #>                                                 
   504       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))] ''simplify_System_parenthesized'' )) #>      
   505       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))] ''isolate_bdvs'' )) #>                   
   506       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))] ''simplify_System'' ))
   507       ) e_2;                 
   508     e__s = Take [e_1, e_2]                                                       
   509   in
   510     Try (Rewrite_Set ''order_system'' ) e__s)                              "
   511 
   512 method met_eqsys_topdown_2x2 : "EqSystem/top_down_substitution/2x2" =
   513   \<open>{rew_ord'="ord_simplify_System", rls' = Rule_Set.Empty, calc = [], 
   514     srls = Rule_Set.append_rules "srls_top_down_2x2" Rule_Set.empty
   515         [\<^rule_thm>\<open>hd_thm\<close>,
   516           \<^rule_thm>\<open>tl_Cons\<close>,
   517           \<^rule_thm>\<open>tl_Nil\<close>], 
   518     prls = prls_triangular, crls = Rule_Set.Empty, errpats = [], nrls = Rule_Set.Empty}\<close>
   519   Program: solve_system.simps
   520   Given: "equalities e_s" "solveForVars v_s"
   521   Where:
   522     "(tl v_s) from v_s occur_exactly_in (NTH 1 (e_s::bool list))"
   523     "    v_s  from v_s occur_exactly_in (NTH 2 (e_s::bool list))"
   524   Find: "solution ss'''"
   525 
   526 method met_eqsys_norm : "EqSystem/normalise" =
   527   \<open>{rew_ord'="tless_true", rls' = Rule_Set.Empty, calc = [], srls = Rule_Set.Empty, prls = Rule_Set.Empty, crls = Rule_Set.Empty,
   528     errpats = [], nrls = Rule_Set.Empty}\<close>
   529 
   530 partial_function (tailrec) solve_system2 :: "bool list => real list => bool list"
   531   where
   532 "solve_system2 e_s v_s = (
   533   let
   534     e__s = (
   535       (Try (Rewrite_Set ''norm_Rational'' )) #>
   536       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))] ''simplify_System_parenthesized'' )) #>
   537       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))] ''isolate_bdvs'' )) #>
   538       (Try (Rewrite_Set_Inst [(''bdv_1'', hd v_s), (''bdv_2'', hd (tl v_s))] ''simplify_System_parenthesized'' )) #>
   539       (Try (Rewrite_Set ''order_system'' ))
   540       ) e_s
   541   in
   542     SubProblem (''EqSystem'', [''LINEAR'', ''system''], [''no_met''])
   543       [BOOL_LIST e__s, REAL_LIST v_s])"
   544 
   545 method met_eqsys_norm_2x2 : "EqSystem/normalise/2x2" =
   546   \<open>{rew_ord'="tless_true", rls' = Rule_Set.Empty, calc = [], 
   547     srls = Rule_Set.append_rules "srls_normalise_2x2" Rule_Set.empty
   548         [\<^rule_thm>\<open>hd_thm\<close>,
   549           \<^rule_thm>\<open>tl_Cons\<close>,
   550           \<^rule_thm>\<open>tl_Nil\<close>], 
   551     prls = Rule_Set.Empty, crls = Rule_Set.Empty, errpats = [], nrls = Rule_Set.Empty}\<close>
   552   Program: solve_system2.simps
   553   Given: "equalities e_s" "solveForVars v_s"
   554   Find: "solution ss'''"
   555 
   556 partial_function (tailrec) solve_system3 :: "bool list => real list => bool list"
   557   where
   558 "solve_system3 e_s v_s = (
   559   let
   560     e__s = (
   561       (Try (Rewrite_Set ''norm_Rational'' )) #>
   562       (Repeat (Rewrite ''commute_0_equality'' )) #>
   563       (Try (Rewrite_Set_Inst [(''bdv_1'', NTH 1 v_s), (''bdv_2'', NTH 2 v_s ),
   564         (''bdv_3'', NTH 3 v_s), (''bdv_3'', NTH 4 v_s )] ''simplify_System_parenthesized'' )) #>
   565       (Try (Rewrite_Set_Inst [(''bdv_1'', NTH 1 v_s), (''bdv_2'', NTH 2 v_s ),
   566         (''bdv_3'', NTH 3 v_s), (''bdv_3'', NTH 4 v_s )] ''isolate_bdvs_4x4'' )) #>
   567       (Try (Rewrite_Set_Inst [(''bdv_1'', NTH 1 v_s), (''bdv_2'', NTH 2 v_s ),
   568         (''bdv_3'', NTH 3 v_s), (''bdv_3'', NTH 4 v_s )] ''simplify_System_parenthesized'' )) #>
   569       (Try (Rewrite_Set ''order_system''))
   570       )  e_s
   571   in
   572     SubProblem (''EqSystem'', [''LINEAR'', ''system''], [''no_met''])
   573       [BOOL_LIST e__s, REAL_LIST v_s])"
   574 
   575 method met_eqsys_norm_4x4 : "EqSystem/normalise/4x4" =
   576   \<open>{rew_ord'="tless_true", rls' = Rule_Set.Empty, calc = [],
   577     srls =
   578       Rule_Set.append_rules "srls_normalise_4x4" srls
   579         [\<^rule_thm>\<open>hd_thm\<close>, \<^rule_thm>\<open>tl_Cons\<close>, \<^rule_thm>\<open>tl_Nil\<close>],
   580     prls = Rule_Set.Empty, crls = Rule_Set.Empty, errpats = [], nrls = Rule_Set.Empty}\<close>
   581   Program: solve_system3.simps
   582   Given: "equalities e_s" "solveForVars v_s"
   583   Find: "solution ss'''"
   584   (*STOPPED.WN06? met ["EqSystem", "normalise", "4x4"] #>#>#>#>#>#>#>#>#>#>#>#>#>@*)
   585 
   586 partial_function (tailrec) solve_system4 :: "bool list => real list => bool list"
   587   where
   588 "solve_system4 e_s v_s = (
   589   let
   590     e_1 = NTH 1 e_s;
   591     e_2 = Take (NTH 2 e_s);
   592     e_2 = (
   593       (Substitute [e_1]) #>
   594       (Try (Rewrite_Set_Inst [(''bdv_1'',NTH 1 v_s), (''bdv_2'',NTH 2 v_s),
   595         (''bdv_3'',NTH 3 v_s), (''bdv_4'',NTH 4 v_s)] ''simplify_System_parenthesized'' )) #>
   596       (Try (Rewrite_Set_Inst [(''bdv_1'',NTH 1 v_s), (''bdv_2'',NTH 2 v_s),
   597         (''bdv_3'',NTH 3 v_s), (''bdv_4'',NTH 4 v_s)] ''isolate_bdvs'' )) #>
   598       (Try (Rewrite_Set_Inst [(''bdv_1'',NTH 1 v_s), (''bdv_2'',NTH 2 v_s),
   599         (''bdv_3'',NTH 3 v_s), (''bdv_4'',NTH 4 v_s)] ''norm_Rational'' ))
   600       ) e_2
   601   in
   602     [e_1, e_2, NTH 3 e_s, NTH 4 e_s])"
   603 
   604 method met_eqsys_topdown_4x4 : "EqSystem/top_down_substitution/4x4" =
   605   \<open>{rew_ord'="ord_simplify_System", rls' = Rule_Set.Empty, calc = [], 
   606     srls = Rule_Set.append_rules "srls_top_down_4x4" srls [], 
   607     prls = Rule_Set.append_rules "prls_tri_4x4_lin_sys" prls_triangular
   608       [\<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in "")], 
   609     crls = Rule_Set.Empty, errpats = [], nrls = Rule_Set.Empty}\<close>
   610   (*FIXXXXME.WN060916: this script works ONLY for exp 7.79 #>#>#>#>#>#>#>#>#>#>*)
   611   Program: solve_system4.simps
   612   Given: "equalities e_s" "solveForVars v_s"
   613   Where: (*accepts missing variables up to diagonal form*)
   614     "(NTH 1 (v_s::real list)) occurs_in (NTH 1 (e_s::bool list))"
   615     "(NTH 2 (v_s::real list)) occurs_in (NTH 2 (e_s::bool list))"
   616     "(NTH 3 (v_s::real list)) occurs_in (NTH 3 (e_s::bool list))"
   617     "(NTH 4 (v_s::real list)) occurs_in (NTH 4 (e_s::bool list))"
   618   Find: "solution ss'''"
   619 
   620 ML \<open>
   621 \<close> ML \<open>
   622 \<close> ML \<open>
   623 \<close>
   624 end