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