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