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