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