test/Tools/isac/Test_Theory.thy
author wneuper <Walther.Neuper@jku.at>
Fri, 18 Aug 2023 18:51:18 +0200
changeset 60733 4097c1317986
parent 60732 18b933a12ab8
child 60734 a3aaca90af25
permissions -rw-r--r--
prepare 16: delete old code 2, shift feedback_to_string Pre_Conds -> I_Model
wneuper@59119
     1
(* use this theory for tests before Build_Isac.thy has succeeded *)
Walther@60576
     2
theory Test_Theory imports "$ISABELLE_ISAC/Build_Isac"
neuper@52102
     3
begin                                                                            
wenzelm@60192
     4
ML_file "$ISABELLE_ISAC/BaseDefinitions/libraryC.sml"
neuper@52102
     5
wneuper@59472
     6
section \<open>code for copy & paste ===============================================================\<close>
wneuper@59472
     7
ML \<open>
Walther@60710
     8
\<close> ML \<open>
Walther@60710
     9
\<close> ML \<open>
walther@59686
    10
"~~~~~ fun xxx , args:"; val () = ();
walther@59686
    11
"~~~~~ and xxx , args:"; val () = ();
Walther@60576
    12
"~~~~~ from fun xxx \<longrightarrow>fun yyy \<longrightarrow>fun zzz , return:"; val () = ();
Walther@60576
    13
"~~~~~ continue fun xxx"; val () = ();
Walther@60729
    14
(*if*) (*then*); (*else*); (*andalso*)  (*case*) (*of*);  (*return value*); (*in*) (*end*);
walther@59686
    15
"xx"
Walther@60629
    16
^ "xxx"   (*+*) (*+++*) (*keep for continuing YYYYY*) (*isa*) (*isa2*) (**)
Walther@60729
    17
\<close> ML \<open> (*//----------- build fun XXXXX -----------------------------------------------------\\*)
Walther@60729
    18
(*//------------------ build fun XXXXX -----------------------------------------------------\\*)
Walther@60729
    19
(*\\------------------ build fun XXXXX -----------------------------------------------------//*)
Walther@60729
    20
\<close> ML \<open> (*\\----------- build fun XXXXX -----------------------------------------------------//*)
Walther@60682
    21
val return_XXXXX = "XXXXX"
Walther@60576
    22
\<close> ML \<open> (*//----------- step into XXXXX -----------------------------------------------------\\*)
Walther@60576
    23
(*//------------------ step into XXXXX -----------------------------------------------------\\*)
Walther@60715
    24
\<close> ML \<open> (*||----------- contine XXXXX ---------------------------------------------------------*)
Walther@60715
    25
(*||------------------ contine XXXXX ---------------------------------------------------------*)
Walther@60576
    26
(*\\------------------ step into XXXXX -----------------------------------------------------//*)
Walther@60576
    27
\<close> ML \<open> (*\\----------- step into XXXXX -----------------------------------------------------//*)
Walther@60710
    28
val "XXXXX" = return_XXXXX;
Walther@60576
    29
Walther@60576
    30
(* final test ... ----------------------------------------------------------------------------*)
Walther@60576
    31
Walther@60576
    32
\<close> ML \<open> (*//----------- inserted hidden code ------------------------------------------------\\*)
Walther@60576
    33
(*//------------------ inserted hidden code ------------------------------------------------\\*)
Walther@60576
    34
(*\\------------------ inserted hidden code ------------------------------------------------//*)
Walther@60576
    35
\<close> ML \<open> (*\\----------- inserted hidden code ------------------------------------------------//*)
Walther@60576
    36
wneuper@59472
    37
\<close>
walther@59723
    38
ML \<open>
walther@59723
    39
\<close> ML \<open>
walther@59723
    40
\<close> ML \<open>
walther@59723
    41
\<close>
wneuper@59472
    42
text \<open>
neuper@52102
    43
  declare [[show_types]] 
neuper@52102
    44
  declare [[show_sorts]]
neuper@52102
    45
  find_theorems "?a <= ?a"
neuper@52102
    46
  
Walther@60576
    47
  print_theorems
neuper@52102
    48
  print_facts
neuper@52102
    49
  print_statement ""
neuper@52102
    50
  print_theory
Walther@60576
    51
  ML_command \<open>Pretty.writeln prt\<close>
Walther@60576
    52
  declare [[ML_print_depth = 999]]
Walther@60576
    53
  declare [[ML_exception_trace]]
wneuper@59472
    54
\<close>
neuper@52102
    55
Walther@60658
    56
section \<open>===================================================================================\<close>
Walther@60725
    57
section \<open>=====  ============================================================================\<close>
wneuper@59472
    58
ML \<open>
wneuper@59472
    59
\<close> ML \<open>
Walther@60733
    60
Walther@60733
    61
\<close> ML \<open>
Walther@60733
    62
\<close>
Walther@60733
    63
Walther@60733
    64
section \<open>===================================================================================\<close>
Walther@60733
    65
section \<open>=====  ============================================================================\<close>
Walther@60733
    66
ML \<open>
Walther@60733
    67
\<close> ML \<open>
Walther@60733
    68
wneuper@59472
    69
\<close> ML \<open>
wneuper@59472
    70
\<close>
walther@60152
    71
Walther@60732
    72
section \<open>===================================================================================\<close>
Walther@60732
    73
section \<open>=====maxi 150a-add-given-Maximum.sml ==============================================\<close>
Walther@60710
    74
ML \<open>
Walther@60710
    75
\<close> ML \<open>
Walther@60715
    76
(* Title:  "Minisubpbl/150a-add-given-Maximum.sml"
Walther@60715
    77
   Author: Walther Neuper 1105
Walther@60715
    78
   (c) copyright due to lincense terms.
Walther@60715
    79
Walther@60715
    80
ATTENTION: the file creates buffer overflow if copied in one piece !
Walther@60715
    81
Walther@60715
    82
Note: This test --- steps into me --- more than once, to a somewhat extreme extent;
Walther@60715
    83
  in order not to get lost while working in Test_Some etc, 
Walther@60715
    84
  re-introduce  ML (*--- step into XXXXX ---*) and Co.
Walther@60715
    85
  and use Sidekick for orientation.
Walther@60715
    86
  Nesting is indicated by  /---   //--   ///-  at the left margin of the comments.
Walther@60715
    87
*)
Walther@60715
    88
Walther@60715
    89
open Ctree;
Walther@60715
    90
open Pos;
Walther@60715
    91
open TermC;
Walther@60715
    92
open Istate;
Walther@60715
    93
open Tactic;
Walther@60715
    94
open Pre_Conds;
Walther@60715
    95
open I_Model;
Walther@60715
    96
open P_Model
Walther@60715
    97
open Rewrite;
Walther@60715
    98
open Step;
Walther@60715
    99
open LItool;
Walther@60715
   100
open LI;
Walther@60715
   101
open Test_Code
Walther@60715
   102
open Specify
Walther@60715
   103
open ME_Misc
Walther@60715
   104
Walther@60715
   105
val (_(*example text*),
Walther@60715
   106
  (model as ("Constants [r = (7::real)]" :: "Maximum A" :: "AdditionalValues [u, v]" :: 
Walther@60715
   107
     "Extremum (A = 2 * u * v - u \<up> 2)" :: 
Walther@60715
   108
     "SideConditions [((u::real) / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]" :: 
Walther@60715
   109
     "SideConditions [((u::real) / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]" :: 
Walther@60715
   110
     "SideConditions [(u::real) / 2 = r * sin \<alpha>, 2 / v = r * cos \<alpha>]" ::
Walther@60722
   111
(*---------------------------------------------,(v::real) / 2 = r * cos \<alpha>]" --- ERROR in example*)
Walther@60715
   112
     "FunctionVariable a" :: "FunctionVariable b" :: "FunctionVariable \<alpha>" :: 
Walther@60715
   113
     "Domain {0 <..< r}" :: "Domain {0 <..< r}" :: "Domain {0 <..< \<pi> / 2}" ::
Walther@60715
   114
     "ErrorBound (\<epsilon> = (0::real))" :: []), 
Walther@60715
   115
   refs as ("Diff_App", 
Walther@60715
   116
     ["univariate_calculus", "Optimisation"],
Walther@60715
   117
     ["Optimisation", "by_univariate_calculus"])))
Walther@60715
   118
  = Store.get (Know_Store.get_expls @{theory}) ["Diff_App-No.123a"] ["Diff_App-No.123a"];
Walther@60715
   119
Walther@60715
   120
\<close> ML \<open>
Walther@60715
   121
val c = [];
Walther@60723
   122
val return_init_calc = 
Walther@60723
   123
 Test_Code.init_calc @{context} [(model, refs)]; (*val Model_Problem = nxt;*)
Walther@60723
   124
\<close> ML \<open> (*/------------ step into init_calc -------------------------------------------------\\*)
Walther@60723
   125
(*/------------------- step into init_calc -------------------------------------------------\\*)
Walther@60723
   126
"~~~~~ fun init_calc , args:"; val (ctxt, [(model, refs as (thy_id, _, _))]) =
Walther@60723
   127
  (@{context}, [(model, refs)]);
Walther@60723
   128
\<close> ML \<open>
Walther@60723
   129
    val thy = thy_id |> ThyC.get_theory ctxt
Walther@60723
   130
\<close> ML \<open>
Walther@60723
   131
    val ((pt, p), tacis) = Step_Specify.initialise' thy (model, refs)
Walther@60723
   132
\<close> ML \<open>
Walther@60723
   133
	  val tac = case tacis of [] => Tactic.Empty_Tac | _ => (#1 o hd) tacis
Walther@60723
   134
\<close> ML \<open>
Walther@60723
   135
	  val f = 
Walther@60728
   136
           TESTg_form ctxt (pt,p);
Walther@60723
   137
\<close> ML \<open>
Walther@60723
   138
"~~~~~ fun TESTg_form , args:"; val (ctxt, ptp) = (ctxt, (pt,p));
Walther@60723
   139
\<close> ML \<open>
Walther@60723
   140
    val (form, _, _) =
Walther@60723
   141
   ME_Misc.pt_extract ctxt ptp
Walther@60723
   142
\<close> ML \<open>
Walther@60723
   143
"~~~~~ fun pt_extract , args:"; val (ctxt, (pt, (p, p_ as Pbl(*Frm,Pbl*)))) = (ctxt, ptp);
Walther@60723
   144
\<close> ML \<open>
Walther@60723
   145
        val ppobj = Ctree.get_obj I pt p
Walther@60723
   146
\<close> ML \<open>
Walther@60728
   147
        val f = if Ctree.is_pblobj ppobj then pt_model ppobj p_ else Ctree.get_obj pt_form pt p;
Walther@60723
   148
\<close> ML \<open>
Walther@60723
   149
            (*if*) Ctree.is_pblobj ppobj (*then*);
Walther@60723
   150
\<close> ML \<open>
Walther@60723
   151
           pt_model ppobj p_ ;
Walther@60723
   152
\<close> ML \<open> (*new...*)
Walther@60723
   153
"~~~~~ fun pt_model , args:"; val ((Ctree.PblObj {probl, spec, origin = (_, spec', hdl), ctxt, ...}), Pbl_) =
Walther@60723
   154
  (ppobj, p_);
Walther@60723
   155
\<close> ML \<open>
Walther@60728
   156
      val (_, pI, _) = Ctree.get_somespec' spec spec';
Walther@60723
   157
\<close> ML \<open>
Walther@60723
   158
(*    val where_ = if pI = Problem.id_empty then []*)
Walther@60723
   159
\<close> ML \<open>
Walther@60723
   160
               (*if*) pI = Problem.id_empty (*else*);
Walther@60723
   161
\<close> ML \<open>
Walther@60723
   162
	          val {where_rls, where_, model, ...} = Problem.from_store ctxt pI
Walther@60723
   163
\<close> ML \<open>
Walther@60723
   164
	          val (_, where_) = 
Walther@60723
   165
 Pre_Conds.check_OLD ctxt where_rls where_
Walther@60728
   166
              (model, I_Model.OLD_to_TEST probl);
Walther@60723
   167
\<close> ML \<open>
Walther@60723
   168
"~~~~~ fun check_OLD , args:"; val (ctxt, where_rls, pres, (model_patt, i_model)) =
Walther@60723
   169
  (ctxt, where_rls, where_, (model, I_Model.OLD_to_TEST probl));
Walther@60723
   170
\<close> ML \<open>
Walther@60729
   171
      val (_, (** )_( **)env_model(**), (env_subst, env_eval)) =
Walther@60723
   172
           of_max_variant model_patt i_model;
Walther@60729
   173
\<close> ML \<open>
Walther@60729
   174
(*maxi*)val "[]" = env_model |> Subst.to_string @{context}
Walther@60729
   175
(*maxi*)val "[]" = env_subst |> Subst.to_string @{context}
Walther@60729
   176
\<close> ML \<open>
Walther@60723
   177
"~~~~~ fun of_max_variant , args:"; val (_, []) = (model_patt, i_model);
Walther@60723
   178
(*\------------------- step into init_calc -------------------------------------------------//*)
Walther@60723
   179
\<close> ML \<open> (*\------------ step into init_calc -------------------------------------------------//*)
Walther@60723
   180
\<close> ML \<open>
Walther@60723
   181
val (p,_,f,nxt,_,pt) = return_init_calc;
Walther@60715
   182
Walther@60715
   183
\<close> ML \<open>
Walther@60715
   184
(*+*)val PblObj {ctxt, probl, ...} = get_obj I pt [];
Walther@60715
   185
(*+*)Proof_Context.theory_of ctxt (*= {Pure, .., Diff_App}*);
Walther@60715
   186
(*+*)val Free ("r", Type ("Real.real", [])) = Syntax.read_term ctxt "r"
Walther@60715
   187
(*+*)val [] = probl
Walther@60715
   188
Walther@60723
   189
\<close> ML \<open> (*ERROR penv_value: no values in 'Constants*)
Walther@60715
   190
val return_me_Model_Problem = 
Walther@60715
   191
           me nxt p c pt; val Add_Given "Constants [r = 7]" = #4 return_me_Model_Problem;
Walther@60715
   192
\<close> ML \<open> (*/------------ step into me Model_Problem ------------------------------------------\\*)
Walther@60715
   193
(*/------------------- step into me Model_Problem ------------------------------------------\\*)
Walther@60715
   194
"~~~~~ fun me , args:"; val (tac, (p:Pos.pos'), (_:Test_Code.NEW), (pt:Ctree.ctree)) = (nxt, p, c, pt);
Walther@60715
   195
      val ctxt = Ctree.get_ctxt pt p
Walther@60715
   196
val return_by_tactic = case
Walther@60715
   197
      Step.by_tactic tac (pt,p) of
Walther@60715
   198
		    ("ok", (_, _, ptp)) => ptp;
Walther@60715
   199
\<close> ML \<open> (*//----------- step into by_tactic -------------------------------------------------\\*)
Walther@60715
   200
(*//------------------ step into by_tactic -------------------------------------------------\\*)
Walther@60715
   201
"~~~~~ fun by_tactic , args:"; val (tac, (ptp as (pt, p))) = (tac, (pt,p));
Walther@60715
   202
val Applicable.Yes tac' = (*case*)
Walther@60715
   203
      Step.check tac (pt, p) (*of*);
Walther@60715
   204
(*+*)val Model_Problem' _ = tac';
Walther@60715
   205
"~~~~~ fun check , args:"; val (tac, (ctree, pos)) = (tac, (pt, p));
Walther@60715
   206
  (*if*) Tactic.for_specify tac (*then*);
Walther@60715
   207
Walther@60715
   208
Specify_Step.check tac (ctree, pos);
Walther@60715
   209
"~~~~~ fun check , args:"; val (Tactic.Model_Problem, (pt, pos as (p, _))) =
Walther@60715
   210
  (tac, (ctree, pos));
Walther@60715
   211
        val (o_model, pI', ctxt) = case Ctree.get_obj I pt p of
Walther@60715
   212
          Ctree.PblObj {origin = (o_model, (_, pI', _), _), ctxt, ...} => (o_model, pI', ctxt)
Walther@60715
   213
	      val {model = model_patt, ...} = Problem.from_store (Ctree.get_ctxt pt pos) pI'
Walther@60715
   214
	      val pbl = I_Model.init_TEST o_model model_patt;
Walther@60715
   215
Walther@60715
   216
val return_check =
Walther@60715
   217
    Applicable.Yes (Tactic.Model_Problem' (pI', I_Model.TEST_to_OLD pbl, []));
Walther@60715
   218
(*\\------------------ step into by_tactic -------------------------------------------------//*)
Walther@60715
   219
\<close> ML \<open> (*\\----------- step into by_tactic -------------------------------------------------//*)
Walther@60715
   220
val (pt, p) = return_by_tactic;
Walther@60715
   221
Walther@60723
   222
\<close> ML \<open>
Walther@60715
   223
val return_do_next = (*case*)
Walther@60715
   224
      Step.do_next p ((pt, Pos.e_pos'), []) (*of*);
Walther@60715
   225
\<close> ML \<open> (*//----------- step into do_next ---------------------------------------------------\\*)
Walther@60715
   226
(*//------------------ step into do_next ---------------------------------------------------\\*)
Walther@60715
   227
"~~~~~ fun do_next , args:"; val (((ip as (_,p_)):pos'), ((ptp as (pt,p), tacis):Calc.state_pre)) =
Walther@60715
   228
  (p, ((pt, e_pos'),[]));
Walther@60729
   229
\<close> ML \<open>
Walther@60729
   230
(*+*)val Pbl = p_
Walther@60729
   231
\<close> ML \<open>
Walther@60715
   232
    val pIopt = get_pblID (pt,ip);
Walther@60715
   233
    (*if*) ip = ([],Res); (* = false*)
Walther@60715
   234
      val _ = (*case*) tacis (*of*);
Walther@60715
   235
      val SOME _ = (*case*) pIopt (*of*);
Walther@60715
   236
Walther@60723
   237
\<close> ML \<open>
Walther@60715
   238
    val ("ok", ([(Add_Given "Constants [r = 7]", _, _)], [], _)) =
Walther@60715
   239
      Step.switch_specify_solve p_ (pt, ip);
Walther@60715
   240
\<close> ML \<open>
Walther@60715
   241
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60729
   242
\<close> ML \<open>
Walther@60715
   243
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60715
   244
Walther@60723
   245
\<close> ML \<open>
Walther@60715
   246
    val ("ok", ([(Add_Given "Constants [r = 7]", _, _)], [], _)) =
Walther@60715
   247
      Step.specify_do_next (pt, input_pos);
Walther@60715
   248
\<close> ML \<open> (*///---------- step into specify_do_next -------------------------------------------\\*)
Walther@60715
   249
(*///----------------- step into specify_do_next -------------------------------------------\\*)
Walther@60715
   250
"~~~~~ fun specify_do_next , args:"; val (ptp as (pt, (p, p_))) = (pt, input_pos);
Walther@60715
   251
Walther@60723
   252
\<close> ML \<open>
Walther@60715
   253
(*  val (_, (p_', tac)) =*)
Walther@60715
   254
val return_find_next_step = (*keep for continuing specify_do_next*)
Walther@60715
   255
   Specify.find_next_step ptp;
Walther@60715
   256
\<close> ML \<open> (*////--------- step into find_next_step --------------------------------------------\\*)
Walther@60715
   257
(*////---------------- step into find_next_step --------------------------------------------\\*)
Walther@60715
   258
"~~~~~ fun find_next_step , args:"; val ((pt, pos as (_, p_))) = (ptp);
Walther@60715
   259
    val {meth = met, origin = origin as (oris, o_refs as (_, pI', mI'), _), probl = pbl,
Walther@60715
   260
      spec = refs, ...} = Calc.specify_data (pt, pos);
Walther@60715
   261
    val ctxt = Ctree.get_ctxt pt pos;
Walther@60715
   262
      (*if*) Ctree.just_created (pt, pos) andalso origin <> Ctree.e_origin (*else*);
Walther@60715
   263
        (*if*) p_ = Pos.Pbl (*then*);
Walther@60715
   264
Walther@60728
   265
   Specify.for_problem ctxt oris (o_refs, refs) (pbl, met);
Walther@60715
   266
\<close> ML \<open> (*/////-------- step into for_problem -----------------------------------------------\\*)
Walther@60715
   267
(*/////--------------- step into for_problem -----------------------------------------------\\*)
Walther@60715
   268
"~~~~~ fun for_problem , args:"; val (ctxt, oris, ((dI', pI', mI'), (dI, pI, mI)), (pbl, met))
Walther@60715
   269
  = (ctxt, oris, (o_refs, refs), (pbl, met));
Walther@60728
   270
    val cdI = if dI = ThyC.id_empty then dI' else dI;
Walther@60715
   271
    val cpI = if pI = Problem.id_empty then pI' else pI;
Walther@60715
   272
    val cmI = if mI = MethodC.id_empty then mI' else mI;
Walther@60715
   273
    val {model = pbt, where_rls, where_, ...} = Problem.from_store ctxt cpI;
Walther@60728
   274
    val {model = mpc, ...} = MethodC.from_store ctxt cmI;
Walther@60729
   275
Walther@60729
   276
\<close> ML \<open> (*//----------- adhoc inserted fun check_OLD ----------------------------------------\\*)
Walther@60732
   277
(*//------------------ adhoc inserted fun check_OLD ----------------------------------------\\*)
Walther@60729
   278
Walther@60729
   279
(*T_TESTold* )
Walther@60729
   280
fun all_lhs_atoms ts = fold (curry and_) (map (fn t =>
Walther@60729
   281
  if can TermC.lhs t then TermC.is_atom (TermC.lhs t) else false) ts) true
Walther@60729
   282
( *T_TEST*)
Walther@60729
   283
fun all_lhs_atoms ts = fold (curry and_) (map (fn t =>
Walther@60729
   284
  if can TermC.lhs t andalso not (TermC.is_num (TermC.lhs t))
Walther@60729
   285
  then TermC.is_atom (TermC.lhs t)
Walther@60729
   286
  else false) ts) true
Walther@60729
   287
(*T_TESTnew*)
Walther@60729
   288
Walther@60729
   289
(*T_TESTold* )
Walther@60726
   290
fun handle_lists id (descr, ts) =
Walther@60729
   291
  if Model_Pattern.is_list_descr descr
Walther@60729
   292
  then if length ts > 1         (*list with more than 1 element needs to collect by [.., .., ..]*)
Walther@60729
   293
    then [(id, TermC.list2isalist (Model_Pattern.typ_of_element descr) ts)] (*AdditValues [u, v]*)
Walther@60729
   294
    else if TermC.is_atom (Library.the_single ts)
Walther@60729
   295
      then [(id, Library.the_single ts)]                                     (* solutions L, ...*)
Walther@60729
   296
      else [(id, TermC.list2isalist (Model_Pattern.typ_of_element descr) ts)](*Constants [r = 7]*)
Walther@60729
   297
  else [(id, Library.the_single ts)]                           (* solveFor x, Extremum (A = ...)*)
Walther@60729
   298
( *T_TEST***)
Walther@60729
   299
fun handle_lists id (descr, ts) =
Walther@60729
   300
  if Model_Pattern.is_list_descr descr
Walther@60729
   301
  then if length ts > 1         (*list with more than 1 element needs to collect by [.., .., ..]*)
Walther@60729
   302
    then if TermC.is_list (hd ts) (*---broken elementwise input to lists---*)
Walther@60729
   303
      then [(id, TermC.list2isalist (Model_Pattern.typ_of_element descr) (map TermC.the_single ts))]
Walther@60729
   304
      else [(id, TermC.list2isalist (Model_Pattern.typ_of_element descr) ts)] (*AdditValues [u, v]*)
Walther@60729
   305
    else if TermC.is_atom (Library.the_single ts)
Walther@60729
   306
      then [(id, Library.the_single ts)]                                     (* solutions L, ...*)
Walther@60729
   307
      else [(id, TermC.list2isalist (Model_Pattern.typ_of_element descr) ts)](*Constants [r = 7]*)
Walther@60729
   308
  else [(id, Library.the_single ts)]                           (* solveFor x, Extremum (A = ...)*)
Walther@60729
   309
(*T_TESTnew*)
Walther@60726
   310
;
Walther@60728
   311
handle_lists: term -> descriptor * term list -> (term * term) list;
Walther@60728
   312
Walther@60722
   313
fun mk_env_model _ (Model_Def.Cor_TEST ((_, []), _)) = []
Walther@60722
   314
  | mk_env_model id (Model_Def.Cor_TEST ((descr, ts), _)) = 
Walther@60726
   315
(*T_TEST.150* )
Walther@60722
   316
    if Model_Pattern.is_list_descr descr 
Walther@60721
   317
    then [(id, TermC.list2isalist (Model_Pattern.typ_of_element descr) ts)]
Walther@60721
   318
    else [(id, hd ts)]
Walther@60726
   319
( *T_TEST*)
Walther@60726
   320
    handle_lists id (descr, ts)
Walther@60726
   321
(*T_TEST.100*)
Walther@60722
   322
  | mk_env_model _ (Model_Def.Syn_TEST _) = [] (*TODO handle correct list elements*)
Walther@60722
   323
  | mk_env_model _ (Model_Def.Inc_TEST ((_, []), _)) = []
Walther@60722
   324
  | mk_env_model id (Model_Def.Inc_TEST ((descr, ts), _)) = 
Walther@60726
   325
(*T_TEST.150* )
Walther@60722
   326
    if Model_Pattern.is_list_descr descr 
Walther@60721
   327
    then [(id, TermC.list2isalist (Model_Pattern.typ_of_element descr) ts)]
Walther@60721
   328
    else [(id, hd ts)]
Walther@60726
   329
( *T_TEST*)
Walther@60726
   330
    handle_lists id (descr, ts)
Walther@60726
   331
(*T_TEST.100*)
Walther@60722
   332
  | mk_env_model _ (Model_Def.Sup_TEST _) = []
Walther@60715
   333
;
Walther@60728
   334
mk_env_model: term -> Model_Def.i_model_feedback_TEST -> Env.T;
Walther@60728
   335
Walther@60728
   336
val (id, (descr, ts)) = (@{term "v_v'i'::bool list"}, (@{term solutions}, [@{term "L::bool list"}]));
Walther@60728
   337
is_atom (Library.the_single [@{term "r = (7::real)"}]);
Walther@60728
   338
Walther@60721
   339
(*  vvvvvvvvvvvv-- REPLACES mk_env_subst_DEL, THUS THE LATTER IS UNCHANGED*)
Walther@60722
   340
fun make_env_model equal_descr_pairs =
Walther@60715
   341
  map (fn ((_, (_, id)), (_, _, _, _, (feedb, _)))
Walther@60722
   342
        => (mk_env_model id feedb)) equal_descr_pairs
Walther@60715
   343
  |> flat
Walther@60715
   344
;
Walther@60728
   345
make_env_model: (Model_Pattern.single * Model_Def.i_model_single_TEST) list -> Env.T;
Walther@60728
   346
Walther@60722
   347
fun switch_type_TEST descr [] = descr
Walther@60722
   348
  | switch_type_TEST (Free (id_string, _)) ts =
Walther@60724
   349
    Free (id_string, ts |> hd |> TermC.lhs |> type_of)
Walther@60722
   350
  | switch_type_TEST descr _ = raise ERROR ("switch_type_TEST undefined argument " ^
Walther@60722
   351
      quote (UnparseC.term (ContextC.for_ERROR ()) descr))
Walther@60722
   352
;
Walther@60728
   353
switch_type_TEST: descriptor -> term list -> descriptor;
Walther@60728
   354
Walther@60729
   355
(*T_TESTold* )
Walther@60722
   356
fun discern_typ _ (_, []) = []
Walther@60722
   357
  | discern_typ id (descr, ts) =
Walther@60723
   358
(*TODO.md "review (descriptor, ts)" REMOVE--------------------------------------\*)
Walther@60723
   359
    let
Walther@60723
   360
      val ts = if Model_Pattern.is_list_descr descr andalso TermC.is_list (hd ts)
Walther@60723
   361
        then TermC.isalist2list (hd ts)
Walther@60723
   362
        else ts
Walther@60723
   363
    in
Walther@60723
   364
(*TODO.md "review (descriptor, ts)" REMOVE--------------------------------------/*)
Walther@60722
   365
  if Model_Pattern.typ_of_element descr = HOLogic.boolT
Walther@60722
   366
    andalso ts |> map TermC.lhs |> all_atoms
Walther@60722
   367
  then
Walther@60722
   368
    if length ts > 1
Walther@60722
   369
    then raise ERROR "model items of type 'bool' in lists with 'lengt ts > 1' NOT YET IMPLEMENTED"
Walther@60722
   370
    else [((switch_type_TEST id ts, TermC.lhs (hd ts)), 
Walther@60722
   371
           (TermC.lhs (hd ts), TermC.rhs (hd ts)))]
Walther@60722
   372
  else []
Walther@60723
   373
(*TODO.md "review (descriptor, ts)" REMOVE--------------------------------------\*)
Walther@60723
   374
    end
Walther@60723
   375
(*TODO.md "review (descriptor, ts)" REMOVE--------------------------------------/*)
Walther@60729
   376
( *T_TEST*)
Walther@60729
   377
fun discern_typ _ (_, []) = []
Walther@60729
   378
  | discern_typ id (descr, ts) =
Walther@60729
   379
(*TODO.md "review (descriptor, ts)" REMOVE--------------------------------------\*)
Walther@60729
   380
    let
Walther@60729
   381
      val ts = if Model_Pattern.is_list_descr descr
Walther@60729
   382
        then if TermC.is_list (hd ts)
Walther@60729
   383
          then ts |> map TermC.isalist2list |> flat
Walther@60729
   384
          else ts
Walther@60729
   385
        else ts
Walther@60729
   386
    in
Walther@60729
   387
(*TODO.md "review (descriptor, ts)" REMOVE--------------------------------------/*)
Walther@60729
   388
  if Model_Pattern.typ_of_element descr = HOLogic.boolT andalso all_lhs_atoms ts
Walther@60729
   389
  then
Walther@60729
   390
    if length ts > 1
Walther@60729
   391
    then (writeln "model items of type 'bool' in lists with 'lengt ts > 1' NOT YET IMPLEMENTED";
Walther@60729
   392
      [])
Walther@60729
   393
    else [((switch_type_TEST id ts, TermC.lhs (hd ts)), 
Walther@60729
   394
           (TermC.lhs (hd ts), TermC.rhs (hd ts)))]
Walther@60729
   395
  else []
Walther@60729
   396
(*TODO.md "review (descriptor, ts)" REMOVE--------------------------------------\*)
Walther@60729
   397
    end
Walther@60729
   398
(*TODO.md "review (descriptor, ts)" REMOVE--------------------------------------/*)
Walther@60723
   399
;
Walther@60728
   400
discern_typ: term -> descriptor * term list -> ((term * term) * (term * term)) list;
Walther@60729
   401
(*T_TESTnew*)
Walther@60728
   402
Walther@60722
   403
fun discern_feedback id (Model_Def.Cor_TEST ((descr, ts), _)) = discern_typ id (descr, ts)
Walther@60722
   404
  | discern_feedback _ (Model_Def.Syn_TEST _) = [] (*TODO: handle correct elements*)
Walther@60722
   405
  | discern_feedback id (Model_Def.Inc_TEST ((descr, ts), _)) = discern_typ id (descr, ts)
Walther@60722
   406
  | discern_feedback _ (Model_Def.Sup_TEST _) = []
Walther@60722
   407
;
Walther@60728
   408
discern_feedback: term -> I_Model.feedback_TEST -> ((term * term) * (term * term)) list;
Walther@60728
   409
Walther@60728
   410
fun make_envs_preconds equal_givens =
Walther@60722
   411
  map (fn ((_, (_, id)), (_, _, _, _, (feedb, _))) => discern_feedback id feedb) equal_givens
Walther@60722
   412
  |> flat
Walther@60722
   413
;
Walther@60722
   414
make_envs_preconds: (Model_Pattern.single * I_Model.single_TEST) list ->
Walther@60728
   415
    ((term * term) * (term * term)) list;
Walther@60728
   416
Walther@60728
   417
(*T_TESTold* )
Walther@60715
   418
fun of_max_variant model_patt i_model =
Walther@60728
   419
( *T_TEST*)
Walther@60728
   420
fun of_max_variant _ [] = ([], [], ([], []))
Walther@60728
   421
  | of_max_variant model_patt i_model =
Walther@60728
   422
(*T_TESTnew*)
Walther@60715
   423
  let
Walther@60715
   424
    val all_variants =
Walther@60715
   425
        map (fn (_, variants, _, _, _) => variants) i_model
Walther@60715
   426
        |> flat
Walther@60715
   427
        |> distinct op =
Walther@60715
   428
    val variants_separated = map (filter_variants' i_model) all_variants
Walther@60715
   429
    val sums_corr = map (cnt_corrects) variants_separated
Walther@60715
   430
    val sum_variant_s = arrange_args1 sums_corr (1, all_variants)
Walther@60715
   431
    val (_, max_variant) = hd (*..crude decision, up to improvement *)
Walther@60715
   432
      (sort (fn ((i, _), (j, _)) => int_ord (i, j)) sum_variant_s)
Walther@60715
   433
    val i_model_max =
Walther@60715
   434
      filter (fn (_, variants, _ , _ ,_) => member (op =) variants max_variant) i_model
Walther@60721
   435
    val equal_descr_pairs = map (get_equal_descr i_model) model_patt |> flat
Walther@60728
   436
(*T_TESTold* )
Walther@60728
   437
    val env_subst = mk_env_subst_DEL equal_descr_pairs
Walther@60728
   438
    val env_eval = mk_env_eval_DEL i_model_max
Walther@60728
   439
  in
Walther@60728
   440
    (i_model_max, env_subst, env_eval)
Walther@60728
   441
  end
Walther@60728
   442
( *T_TEST*)
Walther@60722
   443
    val env_model = make_env_model equal_descr_pairs
Walther@60722
   444
    val equal_givens = filter (fn ((m_field, _), _) => m_field = "#Given") equal_descr_pairs
Walther@60722
   445
    val subst_eval_list = make_envs_preconds equal_givens
Walther@60722
   446
    val (env_subst, env_eval) = split_list subst_eval_list
Walther@60715
   447
  in
Walther@60722
   448
    (i_model_max, env_model, (env_subst, env_eval))
Walther@60715
   449
  end
Walther@60728
   450
(*T_TESTnew*)
Walther@60715
   451
;
Walther@60729
   452
of_max_variant: Model_Pattern.T -> Model_Def.i_model_TEST ->
Walther@60728
   453
       Model_Def.i_model_TEST * Env.T * (env_subst * env_eval);
Walther@60728
   454
Walther@60715
   455
fun check_OLD _ _ [] _  = (true, [])
Walther@60715
   456
  | check_OLD ctxt where_rls pres (model_patt, i_model) =
Walther@60715
   457
    let
Walther@60722
   458
(*T_TESTold* )
Walther@60722
   459
      val (env_subst, env_eval) = sep_variants_envs_OLD model_patt i_model
Walther@60722
   460
( *NEW*)  
Walther@60726
   461
      val (_, env_model, (env_subst, env_eval)) = of_max_variant model_patt i_model
Walther@60715
   462
(*NEW*)
Walther@60715
   463
      val pres_subst = map (TermC.subst_atomic_all env_subst) pres;
Walther@60726
   464
      val pres_subst_other = map (TermC.subst_atomic_all env_model) (map #2 pres_subst);
Walther@60726
   465
      val full_subst = if env_eval = [] then pres_subst_other
Walther@60726
   466
        else map (TermC.subst_atomic_all env_eval) (map #2 pres_subst_other)
Walther@60715
   467
      val evals = map (eval ctxt where_rls) full_subst
Walther@60715
   468
    in
Walther@60726
   469
      (foldl and_ (true, map fst evals), pres_subst_other)
Walther@60723
   470
    end
Walther@60723
   471
;
Walther@60723
   472
check_OLD: Proof.context -> Rule_Def.rule_set -> term list -> Model_Pattern.T * I_Model.T_TEST ->
Walther@60728
   473
  checked;
Walther@60728
   474
Walther@60732
   475
\<close> ML \<open>
Walther@60732
   476
\<close> ML \<open>
Walther@60732
   477
(*T_TESTold*)
Walther@60732
   478
fun check_from_store ctxt where_rls pres env_subst env_eval =
Walther@60727
   479
  let
Walther@60732
   480
    val pres_subst = map (TermC.subst_atomic_all env_subst) pres;
Walther@60732
   481
    val full_subst = map (fn (bool, t) => (bool, subst_atomic env_eval t)) pres_subst
Walther@60732
   482
    val evals = map (eval ctxt where_rls) full_subst
Walther@60732
   483
  in
Walther@60732
   484
    (foldl and_ (true, map fst evals), pres_subst)
Walther@60732
   485
  end
Walther@60732
   486
(*T_TEST*)
Walther@60732
   487
fun check_envs_TEST ctxt where_rls where_ (env_model, (env_subst, env_eval)) =
Walther@60732
   488
  let
Walther@60732
   489
      val pres_subst = map (TermC.subst_atomic_all env_subst) where_;
Walther@60727
   490
      val pres_subst_other = map (TermC.subst_atomic_all env_model) (map #2 pres_subst);
Walther@60727
   491
      val full_subst = if env_eval = [] then pres_subst_other
Walther@60727
   492
        else map (TermC.subst_atomic_all env_eval) (map #2 pres_subst_other)
Walther@60727
   493
      val evals = map (eval ctxt where_rls) full_subst
Walther@60732
   494
  in
Walther@60727
   495
      (foldl and_ (true, map fst evals), pres_subst_other)
Walther@60732
   496
  end
Walther@60727
   497
;
Walther@60732
   498
  check_envs_TEST: Proof.context -> Rule_Set.T -> unchecked -> Env.T * (env_subst * env_eval)
Walther@60732
   499
    -> checked
Walther@60732
   500
\<close> ML \<open>
Walther@60732
   501
(*T_TESTnew*)
Walther@60732
   502
\<close> ML \<open>
Walther@60732
   503
\<close> ML \<open>
Walther@60732
   504
(*\\------------------ adhoc inserted fun check_OLD ----------------------------------------//*)
Walther@60715
   505
\<close> ML \<open> (*\\----------- adhoc inserted fun check_OLD ----------------------------------------//*)
Walther@60722
   506
\<close> ML \<open> (*ERROR mk_env not reasonable for "Inc_TEST Constants [] [__=__, __=__]"*)
Walther@60715
   507
    val return_check_OLD =
Walther@60722
   508
           check_OLD ctxt where_rls where_ (pbt, I_Model.OLD_to_TEST pbl);
Walther@60715
   509
\<close> ML \<open> (*//////------- step into check_OLD -------------------------------------------------\\*)
Walther@60715
   510
(*//////-------------- step into check_OLD -------------------------------------------------\\*)
Walther@60715
   511
"~~~~~ fun check_OLD , args:"; val (ctxt, where_rls, pres, (model_patt, i_model)) =
Walther@60715
   512
  (ctxt, where_rls, where_, (pbt, I_Model.OLD_to_TEST pbl));
Walther@60722
   513
\<close> ML \<open>
Walther@60715
   514
      val return_of_max_variant =
Walther@60728
   515
           of_max_variant model_patt i_model;
Walther@60715
   516
\<close> ML \<open> (*///// //----- step into of_max_variant --------------------------------------------\\*)
Walther@60715
   517
(*///// //------------ step into of_max_variant --------------------------------------------\\*)
Walther@60715
   518
"~~~~~ fun of_max_variant , args:"; val (model_patt, i_model) =
Walther@60715
   519
  (model_patt, i_model);
Walther@60715
   520
Walther@60724
   521
(*with OLD code* )
Walther@60723
   522
(*+*)val "[\n(0, [], false ,#Given, (Inc_TEST Constants [__=__, __=__], Position.T)), \n(0, [], false ,#Find, (Inc_TEST Maximum __, Position.T)), \n(0, [], false ,#Find, (Inc_TEST AdditionalValues [__, __], Position.T)), \n(0, [], false ,#Relate, (Inc_TEST Extremum (__=__), Position.T)), \n(0, [], false ,#Relate, (Inc_TEST SideConditions [__=__, __=__], Position.T))]"
Walther@60723
   523
  = i_model |> I_Model.to_string_TEST @ {context}
Walther@60724
   524
( *with OLD code*)
Walther@60715
   525
\<close> ML \<open> (*with NEW code*)
Walther@60733
   526
(*+*)val "[\n(1, [1, 2, 3], false ,#Given, (Inc_TEST Constants [] [__=__, __=__], Position.T)), \n(2, [1, 2, 3], false ,#Find, (Inc_TEST Maximum __, Position.T)), \n(3, [1, 2, 3], false ,#Find, (Inc_TEST AdditionalValues [] [__, __], Position.T)), \n(4, [1, 2, 3], false ,#Relate, (Inc_TEST Extremum (__=__), Position.T)), \n(5, [1, 2], false ,#Relate, (Inc_TEST SideConditions [] [__=__, __=__], Position.T))]"
Walther@60733
   527
 = i_model |> I_Model.to_string_TEST @{context}
Walther@60715
   528
\<close> ML \<open>
Walther@60715
   529
    val all_variants =
Walther@60715
   530
        map (fn (_, variants, _, _, _) => variants) i_model
Walther@60715
   531
        |> flat
Walther@60715
   532
        |> distinct op =
Walther@60715
   533
    val variants_separated = map (filter_variants' i_model) all_variants
Walther@60715
   534
    val sums_corr = map (cnt_corrects) variants_separated
Walther@60715
   535
    val sum_variant_s = arrange_args1 sums_corr (1, all_variants)
Walther@60715
   536
\<close> ML \<open> (*with NEW code*)
Walther@60715
   537
(*+*)val [(0, 1), (0, 2), (0, 3)] = sum_variant_s
Walther@60722
   538
\<close> ML \<open>
Walther@60715
   539
    val (_, max_variant) = hd (*..crude decision, up to improvement *)
Walther@60715
   540
      (sort (fn ((i, _), (j, _)) => int_ord (i, j)) sum_variant_s)
Walther@60715
   541
    val i_model_max =
Walther@60715
   542
      filter (fn (_, variants, _ , _ ,_) => member (op =) variants max_variant) i_model
Walther@60722
   543
    val equal_descr_pairs = map (get_equal_descr i_model) model_patt |> flat
Walther@60722
   544
\<close> ML \<open>
Walther@60722
   545
(*for building make_env_s -------------------------------------------------------------------\*)
Walther@60722
   546
(*!!!*) val ("#Given", (descr, term), pos) =
Walther@60722
   547
  Model_Pattern.split_descriptor ctxt ("#Given", @{term "Constants [r = (7::real)]"}, Position.none)
Walther@60722
   548
(*!!!*) val patt = equal_descr_pairs |> hd |> #1
Walther@60722
   549
(*!!!*)val equal_descr_pairs =
Walther@60722
   550
  (patt,
Walther@60722
   551
  (1, [1, 2, 3], true, "#Given", (Cor_TEST ((descr, (*!*)TermC.isalist2list(*!*) term), 
Walther@60722
   552
                                                     (TermC.empty, [])), pos)))
Walther@60722
   553
  :: tl equal_descr_pairs
Walther@60722
   554
(*for building make_env_s -------------------------------------------------------------------/*)
Walther@60721
   555
Walther@60722
   556
\<close> ML \<open>
Walther@60728
   557
    val env_model = make_env_model equal_descr_pairs;
Walther@60722
   558
\<close> ML \<open> (*///// ///---- step into make_env_model --------------------------------------------\\*)
Walther@60722
   559
(*///// ///----------- step into make_env_model --------------------------------------------\\*)
Walther@60722
   560
"~~~~~ fun make_env_model , args:"; val (equal_descr_pairs) = (equal_descr_pairs);
Walther@60715
   561
Walther@60715
   562
val xxx = (fn ((_, (_, id)), (_, _, _, _, (feedb, _)))
Walther@60722
   563
       => (mk_env_model id feedb));
Walther@60715
   564
\<close> ML \<open>
Walther@60722
   565
val ((_, (_, id)), (_, _, _, _, (feedb, _))) = nth 4 equal_descr_pairs;
Walther@60722
   566
\<close> ML \<open> (*\\\\\ \\\---- step into make_env_model --------------------------------------------//*)
Walther@60722
   567
(*\\\\\ \\\----------- step into make_env_model --------------------------------------------//*)
Walther@60722
   568
\<close> ML \<open> (*||||| ||----- contine of_max_variant ------------------------------------------------*)
Walther@60722
   569
(*||||| ||------------ contine of_max_variant ------------------------------------------------*)
Walther@60722
   570
Walther@60722
   571
\<close> ML \<open> (*new on 15.7.*)
Walther@60722
   572
    val equal_givens = filter (fn ((m_field, _), _) => m_field = "#Given") equal_descr_pairs
Walther@60728
   573
\<close> ML \<open>
Walther@60722
   574
    val subst_eval_list = make_envs_preconds equal_givens
Walther@60715
   575
\<close> ML \<open>
Walther@60722
   576
val return_make_envs_preconds =
Walther@60722
   577
           make_envs_preconds equal_givens;
Walther@60722
   578
\<close> ML \<open> (*///// ///---- step into make_envs_preconds ----------------------------------------\\*)
Walther@60722
   579
(*///// ///----------- step into make_envs_preconds ----------------------------------------\\*)
Walther@60722
   580
"~~~~~ fun make_envs_preconds , args:"; val (equal_givens) = (equal_givens);
Walther@60722
   581
\<close> ML \<open>
Walther@60722
   582
val xxx = (fn ((_, (_, id)), (_, _, _, _, (feedb, _))) => discern_feedback id feedb)
Walther@60722
   583
;
Walther@60728
   584
xxx: (Model_Pattern.single * I_Model.single_TEST) -> ((term * term) * (term * term)) list;
Walther@60722
   585
\<close> ML \<open>
Walther@60722
   586
val return_discern_feedback =
Walther@60722
   587
           discern_feedback id feedb;
Walther@60722
   588
\<close> ML \<open>
Walther@60722
   589
(*nth 1 equal_descr_pairs* )
Walther@60722
   590
"~~~~~ fun discern_feedback , args:"; val (id, (Model_Def.Cor_TEST ((descr, ts), _))) = (id, feedb);
Walther@60722
   591
( *nth 2 equal_descr_pairs*)
Walther@60722
   592
"~~~~~ fun discern_feedback , args:"; val (id, (Model_Def.Inc_TEST ((descr, ts), _))) = (id, feedb);
Walther@60722
   593
(**)
Walther@60722
   594
\<close> ML \<open>
Walther@60722
   595
\<close> ML \<open>
Walther@60722
   596
\<close> ML \<open>
Walther@60722
   597
(*nth 1 equal_descr_pairs* )
Walther@60722
   598
(*+*)val [((Const ("fixes", typ1), Free ("r", typ2)),
Walther@60722
   599
           (Free ("r", typ3), value))] = return_discern_feedback
Walther@60722
   600
(*+*)val true = typ1 = typ2
Walther@60722
   601
(*+*)val true = typ3 = HOLogic.realT
Walther@60722
   602
(*+*)val "7" = UnparseC.term @{context} value
Walther@60722
   603
( *nth 2 equal_descr_pairs*)
Walther@60728
   604
\<close> ML \<open>
Walther@60722
   605
(*+*)val [] = return_discern_feedback
Walther@60722
   606
(**)
Walther@60722
   607
\<close> ML \<open>
Walther@60722
   608
val return_discern_typ =
Walther@60722
   609
           discern_typ id (descr, ts);
Walther@60722
   610
"~~~~~ fun discern_typ , args:"; val (id, (descr, ts)) = (id, (descr, ts));
Walther@60722
   611
\<close> ML \<open>
Walther@60722
   612
(*nth 1 equal_descr_pairs* )
Walther@60722
   613
(*+*)val [((Const ("fixes", typ1), Free ("r", typ2)),
Walther@60722
   614
           (Free ("r", typ3), value))] = return_discern_typ
Walther@60722
   615
(*+*)val true = typ1 = typ2
Walther@60722
   616
(*+*)val true = typ3 = HOLogic.realT
Walther@60722
   617
(*+*)val "7" = UnparseC.term @{context} value
Walther@60722
   618
( *nth 2 equal_descr_pairs*)
Walther@60728
   619
(*+*)val [] = return_discern_typ;
Walther@60722
   620
(**)
Walther@60722
   621
\<close> ML \<open>
Walther@60722
   622
           switch_type_TEST id ts;
Walther@60722
   623
\<close> ML \<open>
Walther@60722
   624
"~~~~~ fun switch_type_TEST , args:"; val (Const (descr_string, _), ts) = (descr, ts);
Walther@60722
   625
\<close> ML \<open>
Walther@60722
   626
\<close> ML \<open>
Walther@60722
   627
(*nth 1 equal_descr_pairs* )
Walther@60722
   628
val return_switch_type_TEST = Const (descr_string, ts |> hd |> TermC.lhs |> type_of)
Walther@60722
   629
Walther@60722
   630
(*+*)val Const ("Input_Descript.Constants", typ) = return_switch_type_TEST
Walther@60722
   631
(*+*)val Type ("Real.real", []) = typ
Walther@60722
   632
( *nth 2 equal_descr_pairs*)
Walther@60722
   633
(*+*)val return_switch_type_TEST = descr
Walther@60722
   634
(**)
Walther@60722
   635
\<close> ML \<open>
Walther@60722
   636
\<close> ML \<open>
Walther@60722
   637
\<close> ML \<open> (*\\\\\ \\\---- step into make_envs_preconds ----------------------------------------//*)
Walther@60722
   638
(*\\\\\ \\\----------- step into make_envs_preconds ----------------------------------------//*)
Walther@60722
   639
\<close> ML \<open> (*||||| ||----- contine of_max_variant ------------------------------------------------*)
Walther@60722
   640
(*||||| ||------------ contine of_max_variant ------------------------------------------------*)
Walther@60722
   641
    val subst_eval_list = make_envs_preconds equal_givens
Walther@60722
   642
\<close> ML \<open>
Walther@60722
   643
    val (env_subst, env_eval) = split_list subst_eval_list
Walther@60722
   644
\<close> ML \<open>
Walther@60729
   645
(*maxi*)val "[\"\n(fixes, [r = 7])\"]" = env_model |> Subst.to_string @{context}
Walther@60729
   646
(*maxi*)val "[\"\n(fixes, r)\"]" = env_subst |> Subst.to_string @{context}
Walther@60729
   647
\<close> ML \<open>
Walther@60722
   648
val return_of_max_variant = (i_model_max, env_model, (env_subst, env_eval)); (*GOON*)
Walther@60722
   649
(*\\\\\ \\------------ step into of_max_variant --------------------------------------------//*)
Walther@60722
   650
\<close> ML \<open> (*\\\\\ \\----- step into of_max_variant --------------------------------------------//*)
Walther@60722
   651
      val (i_model_max, env_model, (env_subst, env_eval)) = return_of_max_variant
Walther@60729
   652
\<close> ML \<open>
Walther@60722
   653
(*!!!/----- we had a helpful argument for constructing env_model and (env_subst, env_eval)---\*)
Walther@60722
   654
      val (i_max_variant, env_model, (env_subst, env_eval)) = (i_model_max, [], ([], []))
Walther@60722
   655
(*!!!\----- we had a helpful argument for constructing env_model and (env_subst, env_eval)---/*) 
Walther@60722
   656
\<close> ML \<open> (*||||||------- contine check_OLD -----------------------------------------------------*)
Walther@60715
   657
(*||||||-------------- contine check_OLD -----------------------------------------------------*)
Walther@60721
   658
\<close> ML \<open>
Walther@60715
   659
      val pres_subst = map (TermC.subst_atomic_all env_subst) pres;
Walther@60726
   660
      val pres_subst_other = map (TermC.subst_atomic_all env_model) (map #2 pres_subst);
Walther@60726
   661
      val full_subst = if env_eval = [] then pres_subst_other
Walther@60726
   662
        else map (TermC.subst_atomic_all env_eval) (map #2 pres_subst_other)
Walther@60715
   663
      val evals = map (eval ctxt where_rls) full_subst
Walther@60715
   664
val return_ = (i_model_max, env_subst, env_eval)
Walther@60715
   665
(*\\\\\\\------------- step into check_OLD -------------------------------------------------//*)
Walther@60715
   666
\<close> ML \<open> (*\\\\\\------- step into check_OLD -------------------------------------------------//*)
Walther@60723
   667
\<close> ML \<open>
Walther@60715
   668
val (preok, _) = return_check_OLD;
Walther@60715
   669
\<close> ML \<open> (*|||||-------- contine for_problem ---------------------------------------------------*)
Walther@60715
   670
(*|||||--------------- contine for_problem ---------------------------------------------------*)
Walther@60715
   671
    (*if*) dI' = ThyC.id_empty andalso dI = ThyC.id_empty (*else*);
Walther@60715
   672
      (*if*) pI' = Problem.id_empty andalso pI = Problem.id_empty (*else*);
Walther@60715
   673
val NONE =
Walther@60715
   674
     (*case*) find_first (I_Model.is_error o #5) pbl (*of*);
Walther@60715
   675
Walther@60715
   676
        (*case*)
Walther@60715
   677
   Specify.item_to_add (ThyC.get_theory ctxt
Walther@60715
   678
            (if dI = ThyC.id_empty then dI' else dI)) oris pbt pbl (*of*);
Walther@60715
   679
"~~~~~ fun item_to_add , args:"; val (thy, oris, _, itms)
Walther@60715
   680
  = ((ThyC.get_theory ctxt (if dI = ThyC.id_empty then dI' else dI)), oris, pbt, pbl);
Walther@60715
   681
      fun false_and_not_Sup (_, _, false, _, I_Model.Sup _) = false
Walther@60715
   682
        | false_and_not_Sup (_, _, false, _, _) = true
Walther@60715
   683
        | false_and_not_Sup _ = false
Walther@60715
   684
Walther@60715
   685
      val v = if itms = [] then 1 else Pre_Conds.max_variant itms
Walther@60715
   686
      val vors = if v = 0 then oris
Walther@60715
   687
        else filter ((fn variant =>
Walther@60715
   688
            fn (_, variants, m_field, _, _) => member op= variants variant andalso m_field <> "#undef")
Walther@60715
   689
          v) oris
Walther@60715
   690
Walther@60715
   691
(*+*)val "[\n(1, [\"1\", \"2\", \"3\"], #Given, Constants, [\"[r = 7]\"]), \n(2, [\"1\", \"2\", \"3\"], #Find, Maximum, [\"A\"]), \n(3, [\"1\", \"2\", \"3\"], #Find, AdditionalValues, [\"[u]\", \"[v]\"]), \n(4, [\"1\", \"2\", \"3\"], #Relate, Extremum, [\"A = 2 * u * v - u \<up> 2\"]), \n(5, [\"1\", \"2\"], #Relate, SideConditions, [\"[(u / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]\"])]"
Walther@60715
   692
(*+*)  = vors |> O_Model.to_string @{context}
Walther@60715
   693
Walther@60715
   694
      val vits = if v = 0 then itms               (* because of dsc without dat *)
Walther@60715
   695
  	    else filter ((fn variant =>
Walther@60715
   696
            fn (_, variants, _, _, _) => member op= variants variant)
Walther@60715
   697
          v) itms;                                (* itms..vat *)
Walther@60715
   698
Walther@60715
   699
      val icl = filter false_and_not_Sup vits;    (* incomplete *)
Walther@60715
   700
Walther@60715
   701
      (*if*) icl = [] (*else*);
Walther@60715
   702
(*+*)val "[\n(1 ,[1, 2, 3] ,false ,#Given ,Inc Constants [] ,(Constants, [])), \n(2 ,[1, 2, 3] ,false ,#Find ,Inc Maximum ,(Maximum, [])), \n(3 ,[1, 2, 3] ,false ,#Find ,Inc AdditionalValues [] ,(AdditionalValues, [])), \n(4 ,[1, 2, 3] ,false ,#Relate ,Inc Extremum ,(Extremum, [])), \n(5 ,[1, 2] ,false ,#Relate ,Inc SideConditions [] ,(SideConditions, []))]"
Walther@60715
   703
(*+*)  = icl |> I_Model.to_string @{context}
Walther@60715
   704
(*+*)val "(1 ,[1, 2, 3] ,false ,#Given ,Inc Constants [] ,(Constants, []))"
Walther@60715
   705
(*+*)  = hd icl |> I_Model.single_to_string @{context}
Walther@60715
   706
Walther@60715
   707
(*++*)val feedback = (fn (_, _, _, _, feedback) => feedback) (hd icl)
Walther@60715
   708
(*++*)val Const ("Input_Descript.Constants", _) = I_Model.descriptor feedback
Walther@60715
   709
(*++*)val [] = I_Model.o_model_values feedback
Walther@60715
   710
Walther@60715
   711
(*+*)val "[\n(1, [\"1\", \"2\", \"3\"], #Given, Constants, [\"[r = 7]\"]), \n(2, [\"1\", \"2\", \"3\"], #Find, Maximum, [\"A\"]), \n(3, [\"1\", \"2\", \"3\"], #Find, AdditionalValues, [\"[u]\", \"[v]\"]), \n(4, [\"1\", \"2\", \"3\"], #Relate, Extremum, [\"A = 2 * u * v - u \<up> 2\"]), \n(5, [\"1\", \"2\"], #Relate, SideConditions, [\"[(u / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]\"])]"
Walther@60715
   712
(*+*)  = vors |> O_Model.to_string @{context}
Walther@60715
   713
Walther@60715
   714
val SOME ori =
Walther@60715
   715
        (*case*) find_first ((fn (_, _, _, _, feedback) => fn (_, _, _, d, ts) =>
Walther@60715
   716
           d = I_Model.descriptor feedback andalso subset op = (I_Model.o_model_values feedback, ts))
Walther@60715
   717
         (hd icl)) vors (*of*);
Walther@60715
   718
Walther@60715
   719
(*+*)val "(1, [\"1\", \"2\", \"3\"], #Given, Constants, [\"[r = 7]\"])" =
Walther@60715
   720
(*+*)  ori |> O_Model.single_to_string @{context}
Walther@60715
   721
\<close> ML \<open> (*\\\\\-------- step into for_problem -----------------------------------------------//*)
Walther@60715
   722
(*\\\\\--------------- step into for_problem -----------------------------------------------//*)
Walther@60715
   723
\<close> ML \<open> (*\\\\--------- step into find_next_step --------------------------------------------//*)
Walther@60715
   724
(*\\\\---------------- step into find_next_step --------------------------------------------//*)
Walther@60715
   725
\<close> ML \<open> (*|||---------- continuing specify_do_next --------------------------------------------*)
Walther@60715
   726
(*|||----------------- continuing specify_do_next --------------------------------------------*)
Walther@60723
   727
\<close> ML \<open> (*while make Test_Theory work*)
Walther@60715
   728
val (_, (p_', tac)) = return_find_next_step (*kept for continuing specify_do_next*)
Walther@60715
   729
Walther@60715
   730
    val ist_ctxt =  Ctree.get_loc pt (p, p_)
Walther@60715
   731
Walther@60715
   732
(*+*)val Add_Given "Constants [r = 7]" = tac
Walther@60715
   733
val _ =
Walther@60715
   734
    (*case*) tac (*of*);
Walther@60715
   735
Walther@60723
   736
\<close> ML \<open>
Walther@60715
   737
Step_Specify.by_tactic_input tac (pt, (p, p_'));
Walther@60723
   738
\<close> ML \<open>
Walther@60715
   739
"~~~~~ fun by_tactic_input , args:"; val ((Tactic.Add_Given ct), ptp) =
Walther@60715
   740
  (tac, (pt, (p, p_')));
Walther@60715
   741
Walther@60715
   742
   Specify.by_Add_ "#Given" ct ptp;
Walther@60715
   743
"~~~~~ fun by_Add_ , args:"; val (m_field, ct ,(pt, pos as (_, p_))) =
Walther@60715
   744
  ("#Given", ct, ptp);
Walther@60715
   745
    val (met, oris, (_, pI', mI'), pbl, (_, pI, mI), ctxt) = SpecificationC.get_data (pt, pos)
Walther@60715
   746
    val (i_model, m_patt) =
Walther@60715
   747
       if p_ = Pos.Met then
Walther@60715
   748
         (met,
Walther@60715
   749
           (if mI = MethodC.id_empty then mI' else mI) |> MethodC.from_store ctxt |> #model)
Walther@60715
   750
       else
Walther@60715
   751
         (pbl,
Walther@60715
   752
           (if pI = Problem.id_empty then pI' else pI) |> Problem.from_store ctxt |> #model);
Walther@60715
   753
Walther@60715
   754
      (*case*)
Walther@60715
   755
   I_Model.check_single ctxt m_field oris i_model m_patt ct (*of*);
Walther@60715
   756
"~~~~~ fun check_single , args:"; val (ctxt, m_field, o_model, i_model, m_patt, (str(*, pos*))) =
Walther@60715
   757
  (ctxt, m_field, oris, i_model, m_patt, ct);
Walther@60715
   758
        val (t as (descriptor $ _)) = Syntax.read_term ctxt str
Walther@60715
   759
Walther@60723
   760
(*+*)val "Constants [r = 7]" = UnparseC.term @{context} t;
Walther@60715
   761
Walther@60715
   762
        val SOME m_field' =
Walther@60715
   763
          (*case*) Model_Pattern.get_field descriptor m_patt (*of*);
Walther@60715
   764
           (*if*) m_field <> m_field' (*else*);
Walther@60715
   765
Walther@60715
   766
(*+*)val "#Given" = m_field; val "#Given" = m_field'
Walther@60715
   767
Walther@60715
   768
val ("", ori', all) =
Walther@60715
   769
          (*case*) O_Model.contains ctxt m_field o_model t (*of*);
Walther@60715
   770
Walther@60723
   771
\<close> ML \<open>
Walther@60715
   772
(*+*)val (_, _, _, _, vals) = hd o_model;
Walther@60723
   773
\<close> ML \<open>
Walther@60723
   774
(*+*)val "Constants [r = 7]" = UnparseC.term @{context} (@{term Constants} $ (hd vals));
Walther@60715
   775
(*+*)if "[\n(1, [\"1\", \"2\", \"3\"], #Given, Constants, [\"[r = 7]\"]), " ^ 
Walther@60715
   776
(*+*)    "\n(2, [\"1\", \"2\", \"3\"], #Find, Maximum, [\"A\"]), " ^ 
Walther@60715
   777
(*+*)    "\n(3, [\"1\", \"2\", \"3\"], #Find, AdditionalValues, [\"[u]\", \"[v]\"]), " ^ 
Walther@60715
   778
(*+*)    "\n(4, [\"1\", \"2\", \"3\"], #Relate, Extremum, [\"A = 2 * u * v - u \<up> 2\"]), " ^ 
Walther@60715
   779
(*+*)    "\n(5, [\"1\", \"2\"], #Relate, SideConditions, [\"[(u / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]\"]), " ^ 
Walther@60715
   780
(*+*)    "\n(6, [\"3\"], #Relate, SideConditions, [\"[u / 2 = r * sin \<alpha>]\", \"[2 / v = r * cos \<alpha>]\"]), " ^ 
Walther@60715
   781
(*+*)    "\n(7, [\"1\"], #undef, FunctionVariable, [\"a\"]), " ^ 
Walther@60715
   782
(*+*)    "\n(8, [\"2\"], #undef, FunctionVariable, [\"b\"]), \n(9, [\"3\"], #undef, FunctionVariable, [\"\<alpha>\"]), " ^ 
Walther@60715
   783
(*+*)    "\n(10, [\"1\", \"2\"], #undef, Input_Descript.Domain, [\"{0<..<r}\"]), " ^ 
Walther@60715
   784
(*+*)    "\n(11, [\"3\"], #undef, Input_Descript.Domain, [\"{0<..<\<pi> / 2}\"]), " ^ 
Walther@60715
   785
(*+*)    "\n(12, [\"1\", \"2\", \"3\"], #undef, ErrorBound, [\"\<epsilon> = 0\"])]"
Walther@60723
   786
(*+*)= O_Model.to_string @{context} o_model then () else error "o_model CHANGED";
Walther@60723
   787
\<close> ML \<open>
Walther@60715
   788
Walther@60715
   789
  (*case*) is_notyet_input ctxt i_model all ori' m_patt (*of*);
Walther@60715
   790
"~~~~~ fun is_notyet_input , args:"; val (ctxt, itms, all, (i, v, f, d, ts), pbt) =
Walther@60715
   791
  (ctxt, i_model, all, ori', m_patt);
Walther@60715
   792
val SOME (_, (_, pid)) =
Walther@60715
   793
  (*case*) find_first (eq1 d) pbt (*of*);
Walther@60715
   794
(*local*)fun eq3 f d (_, _, _, f', itm_) = f = f' andalso d = (I_Model.descriptor itm_);(*local*)
Walther@60723
   795
\<close> ML \<open>
Walther@60715
   796
val SOME (_, _, _, _, itm_) =
Walther@60715
   797
    (*case*) find_first (eq3 f d) itms (*of*);
Walther@60715
   798
val ts' = inter op = (o_model_values itm_) ts;
Walther@60715
   799
            (*if*) subset op = (ts, ts') (*else*);
Walther@60715
   800
val return_is_notyet_input = ("", 
Walther@60715
   801
           ori_2itm itm_ pid all (i, v, f, d, subtract op = ts' ts));
Walther@60723
   802
\<close> ML \<open>
Walther@60715
   803
"~~~~~ fun ori_2itm , args:"; val (itm_, pid, all, (id, vt, fd, d, ts)) =
Walther@60715
   804
  (itm_, pid, all, (i, v, f, d, subtract op = ts' ts));
Walther@60723
   805
\<close> ML \<open>
Walther@60715
   806
    val ts' = union op = (o_model_values itm_) ts;
Walther@60715
   807
    val pval = [Input_Descript.join'''' (d, ts')]
Walther@60715
   808
    val complete = if eq_set op = (ts', all) then true else false
Walther@60715
   809
Walther@60723
   810
(*+*)val "Inc Constants [] ,(Constants, [])" = itm_ |> I_Model.feedback_to_string @{context}
Walther@60715
   811
\<close> ML \<open> (*\\\---------- step into specify_do_next -------------------------------------------//*)
Walther@60715
   812
(*\\\----------------- step into specify_do_next -------------------------------------------//*)
Walther@60715
   813
\<close> ML \<open> (*\\----------- step into do_next ---------------------------------------------------//*)
Walther@60715
   814
(*\\------------------ step into do_next ---------------------------------------------------//*)
Walther@60726
   815
val ("ok", (ts as (_, _, _) :: _, _, (pt, p))) = return_do_next
Walther@60726
   816
\<close> ML \<open>
Walther@60715
   817
Walther@60715
   818
\<close> ML \<open>(*|------------- continue with me_Model_Problem ----------------------------------------*)
Walther@60715
   819
(*|------------------- continue with me_Model_Problem ----------------------------------------*)
Walther@60726
   820
\<close> ML \<open>
Walther@60715
   821
Walther@60715
   822
val tacis as (_::_) =
Walther@60715
   823
        (*case*) ts (*of*);
Walther@60715
   824
          val (tac, _, _) = last_elem tacis
Walther@60715
   825
Walther@60723
   826
\<close> ML \<open>
Walther@60726
   827
val return_Model_Problem = (p, [] : NEW, TESTg_form ctxt (pt, p), tac, Celem.Sundef, pt);
Walther@60723
   828
\<close> ML \<open>
Walther@60715
   829
\<close> ML \<open> (*//----------- step into TESTg_form ------------------------------------------------\\*)
Walther@60715
   830
(*//------------------ step into TESTg_form ------------------------------------------------\\*)
Walther@60715
   831
"~~~~~ fun TESTg_form , args:"; val (ctxt, ptp) = (ctxt, (pt, p));
Walther@60715
   832
Walther@60723
   833
\<close> ML \<open>
Walther@60715
   834
    val (form, _, _) =
Walther@60715
   835
   ME_Misc.pt_extract ctxt ptp;
Walther@60723
   836
\<close> ML \<open>
Walther@60715
   837
"~~~~~ fun pt_extract , args:"; val (ctxt, (pt, (p, p_(*Frm,Pbl*)))) = (ctxt, ptp);
Walther@60723
   838
\<close> ML \<open>
Walther@60715
   839
        val ppobj = Ctree.get_obj I pt p
Walther@60723
   840
\<close> ML \<open>
Walther@60715
   841
        val f = if Ctree.is_pblobj ppobj then pt_model ppobj p_ else Ctree.get_obj pt_form pt p;
Walther@60723
   842
\<close> ML \<open>
Walther@60715
   843
          (*if*) Ctree.is_pblobj ppobj (*then*);
Walther@60715
   844
Walther@60723
   845
\<close> ML \<open>
Walther@60715
   846
           pt_model ppobj p_;
Walther@60723
   847
\<close> ML \<open>
Walther@60723
   848
"~~~~~ fun pt_model , args:"; val ((Ctree.PblObj {probl, spec, origin = (_, spec', hdl), ctxt, ...}), 
Walther@60723
   849
  Pbl(*Frm,Pbl*)) = (ppobj, p_);
Walther@60723
   850
\<close> ML \<open>
Walther@60728
   851
      val (_, pI, _) = Ctree.get_somespec' spec spec';
Walther@60723
   852
\<close> ML \<open>
Walther@60723
   853
                 (*if*) pI = Problem.id_empty (*else*); 
Walther@60723
   854
(*    val where_ = if pI = Problem.id_empty then []*)
Walther@60723
   855
(*      else                                       *)
Walther@60723
   856
(*        let                                      *)
Walther@60723
   857
	          val {where_rls, where_, model, ...} = Problem.from_store ctxt pI
Walther@60723
   858
\<close> ML \<open>
Walther@60723
   859
	          val (_, where_) = (*Pre_Conds.*)check_OLD ctxt where_rls where_ (model, I_Model.OLD_to_TEST probl)
Walther@60723
   860
(*        in where_ end                           *)
Walther@60723
   861
\<close> ML \<open>  (*ERROR from previous*)
Walther@60715
   862
	    val allcorrect = I_Model.is_complete probl andalso foldl and_ (true, (map #1 where_))
Walther@60723
   863
\<close> ML \<open>
Walther@60715
   864
val return_pt_model = Ctree.ModSpec (allcorrect, Pos.Pbl, hdl, probl, where_, spec)
Walther@60723
   865
\<close> ML \<open>
Walther@60715
   866
\<close> ML \<open>(*|------------- continue with TESTg_form ----------------------------------------------*)
Walther@60715
   867
(*|------------------- continue with TESTg_form ----------------------------------------------*)
Walther@60723
   868
\<close> ML \<open> (*while make Test_Theory work*)
Walther@60715
   869
val Ctree.ModSpec (spec as (_, p_, _, gfr, where_, _)) =
Walther@60715
   870
    (*case*) form (*of*);
Walther@60715
   871
    Test_Out.PpcKF (  (Test_Out.Problem [],
Walther@60715
   872
 			P_Model.from (Proof_Context.theory_of ctxt) gfr where_));
Walther@60715
   873
Walther@60715
   874
   P_Model.from (Proof_Context.theory_of ctxt) gfr where_;
Walther@60715
   875
"~~~~~ fun from , args:"; val (thy, itms, where_) = ((Proof_Context.theory_of ctxt), gfr, where_);
Walther@60715
   876
    fun coll model [] = model
Walther@60715
   877
      | coll model ((_, _, _, field, itm_) :: itms) =
Walther@60715
   878
        coll (add_sel_ppc thy field model (item_from_feedback thy itm_)) itms;
Walther@60715
   879
Walther@60715
   880
 val gfr = coll P_Model.empty itms;
Walther@60715
   881
"~~~~~ fun coll , args:"; val (model, ((aaa, bbb_,ccc_, field, itm_) :: itms))
Walther@60715
   882
  = (P_Model.empty, itms);
Walther@60715
   883
Walther@60715
   884
(*+*)val 4 = length itms;
Walther@60715
   885
(*+*)val itm = nth 1 itms;
Walther@60715
   886
Walther@60715
   887
           coll P_Model.empty [itm];
Walther@60715
   888
"~~~~~ fun coll , iterate:"; val (model, ((aaa, bbb_,ccc_, field, itm_) :: []))
Walther@60715
   889
  = (P_Model.empty, [itm]);
Walther@60715
   890
Walther@60715
   891
          (add_sel_ppc thy field model (item_from_feedback thy itm_));
Walther@60715
   892
"~~~~~ fun add_sel_ppc , args:"; val ((_: theory), sel, {Given = gi, Where = wh, Find = fi, With = wi, Relate = re}, x )
Walther@60715
   893
  = (thy, field, model, (item_from_feedback thy itm_));
Walther@60715
   894
Walther@60715
   895
   P_Model.item_from_feedback thy itm_;
Walther@60715
   896
"~~~~~ fun item_from_feedback , args:"; val (thy, (I_Model.Inc ((d, ts), _))) = (thy, itm_);
Walther@60715
   897
   P_Model.Incompl (UnparseC.term_in_thy thy (Input_Descript.join (d, ts)));
Walther@60715
   898
(*\\------------------ step into TESTg_form ------------------------------------------------//*)
Walther@60715
   899
\<close> ML \<open> (*\\----------- step into TESTg_form ------------------------------------------------//*)
Walther@60715
   900
\<close> ML \<open> (*\------------ step into me Model_Problem ------------------------------------------//*)
Walther@60715
   901
(*\------------------- step into me Model_Problem ------------------------------------------//*)
Walther@60723
   902
\<close> ML \<open> (*while make Test_Theory work*)
Walther@60715
   903
val (p, _, f, nxt, _, pt) = return_me_Model_Problem
Walther@60715
   904
Walther@60715
   905
\<close> ML \<open> (*------------- contine me's ----------------------------------------------------------*)
Walther@60715
   906
(*-------------------- contine me's ----------------------------------------------------------*)
Walther@60723
   907
val return_me_add_find_Constants = me nxt p c pt;
Walther@60723
   908
                                      val Add_Find "Maximum A" = #4 return_me_add_find_Constants;
Walther@60723
   909
\<close> ML \<open> (*/------------ step into me_add_find_Constants -------------------------------------\\*)
Walther@60723
   910
(*/------------------- step into me_add_find_Constants -------------------------------------\\*)
Walther@60723
   911
"~~~~~ fun me , args:"; val (tac as Add_Given "Constants [r = 7]", p, _(*NEW remove*), pt) =
Walther@60723
   912
  (nxt, p, c, pt);
Walther@60723
   913
\<close> ML \<open>
Walther@60723
   914
      val ctxt = Ctree.get_ctxt pt p
Walther@60723
   915
\<close> ML \<open>
Walther@60723
   916
(*+*)val PblObj {probl = (1, [1, 2, 3], false, "#Given", Inc 
Walther@60723
   917
    ((Const ("Input_Descript.Constants", _), []), _)) :: _, ...}  = get_obj I pt (fst p)
Walther@60723
   918
(*-------------------------------------------^^--*)
Walther@60723
   919
val return_step_by_tactic = (*case*) 
Walther@60723
   920
      Step.by_tactic tac (pt, p) (*of*);
Walther@60723
   921
\<close> ML \<open> (*//----------- step into Step.by_tactic --------------------------------------------\\*)
Walther@60723
   922
(*//------------------ step into Step.by_tactic --------------------------------------------\\*)
Walther@60723
   923
"~~~~~ fun by_tactic , args:"; val (tac, (ptp as (pt, p))) = (tac, (pt, p));
Walther@60723
   924
\<close> ML \<open>
Walther@60723
   925
val Applicable.Yes tac' =
Walther@60723
   926
    (*case*) check tac (pt, p) (*of*);
Walther@60723
   927
\<close> ML \<open>
Walther@60723
   928
	    (*if*) Tactic.for_specify' tac' (*then*);
Walther@60723
   929
\<close> ML \<open>
Walther@60723
   930
Step_Specify.by_tactic tac' ptp;
Walther@60723
   931
\<close> ML \<open>
Walther@60723
   932
"~~~~~ fun by_tactic , args:"; val ((Tactic.Add_Given' (ct, _)), (pt, p)) = (tac', ptp);
Walther@60723
   933
\<close> ML \<open>
Walther@60723
   934
   Specify.by_Add_ "#Given" ct (pt, p);
Walther@60723
   935
\<close> ML \<open>
Walther@60723
   936
"~~~~~ fun by_Add_ , args:"; val (m_field, ct, (pt, pos as (_, p_))) = ("#Given", ct, (pt, p));
Walther@60723
   937
\<close> ML \<open>
Walther@60728
   938
    val (met, oris, (_, pI', mI'), pbl, (_, pI, mI), ctxt) = SpecificationC.get_data (pt, pos);
Walther@60723
   939
\<close> ML \<open>
Walther@60723
   940
(*  val (i_model, m_patt) =*)
Walther@60723
   941
       (*if*) p_ = Pos.Met (*else*);
Walther@60723
   942
\<close> ML \<open>
Walther@60723
   943
val return_by_Add_ =
Walther@60723
   944
         (pbl,
Walther@60723
   945
           (if pI = Problem.id_empty then pI' else pI) |> Problem.from_store ctxt |> #model)
Walther@60723
   946
\<close> ML \<open>
Walther@60723
   947
val I_Model.Add i_single =
Walther@60723
   948
      (*case*) I_Model.check_single ctxt m_field oris i_model m_patt ct (*of*);
Walther@60723
   949
\<close> ML \<open>
Walther@60723
   950
\<close> ML \<open>
Walther@60723
   951
	          val i_model' =
Walther@60728
   952
   I_Model.add_single (Proof_Context.theory_of ctxt) i_single i_model;
Walther@60723
   953
\<close> ML \<open>
Walther@60723
   954
"~~~~~ fun add_single , args:"; val (thy, itm, model) =
Walther@60723
   955
  ((Proof_Context.theory_of ctxt), i_single, i_model);
Walther@60723
   956
\<close> ML \<open>
Walther@60723
   957
    fun eq_untouched d (0, _, _, _, itm_) = (d = I_Model.descriptor itm_)
Walther@60723
   958
      | eq_untouched _ _ = false
Walther@60723
   959
\<close> ML \<open>
Walther@60723
   960
    val model' = case I_Model.seek_ppc (#1 itm) model of
Walther@60723
   961
      SOME _ => overwrite_ppc thy itm model (*itm updated in is_notyet_input WN.11.03*)
Walther@60723
   962
\<close> ML \<open>
Walther@60723
   963
\<close> ML \<open>
Walther@60723
   964
\<close> ML \<open>
Walther@60723
   965
\<close> ML \<open> (*||----------- contine Step.by_tactic ------------------------------------------------*)
Walther@60723
   966
(*||------------------ contine Step.by_tactic ------------------------------------------------*)
Walther@60723
   967
(*\\------------------ step into Step.by_tactic --------------------------------------------//*)
Walther@60723
   968
\<close> ML \<open> (*\\----------- step into Step.by_tactic --------------------------------------------//*)
Walther@60723
   969
val ("ok", (_, _, ptp)) = return_step_by_tactic;
Walther@60723
   970
\<close> ML \<open>
Walther@60728
   971
      val (pt, p) = ptp;
Walther@60723
   972
\<close> ML \<open>
Walther@60723
   973
        (*case*) 
Walther@60723
   974
      Step.do_next p ((pt, Pos.e_pos'), []) (*of*);
Walther@60723
   975
\<close> ML \<open>
Walther@60723
   976
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis)) =
Walther@60723
   977
  (p, ((pt, Pos.e_pos'), []));
Walther@60723
   978
\<close> ML \<open>
Walther@60723
   979
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60723
   980
\<close> ML \<open>
Walther@60723
   981
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60723
   982
\<close> ML \<open>
Walther@60723
   983
val _ =
Walther@60723
   984
      (*case*) tacis (*of*);
Walther@60723
   985
\<close> ML \<open>
Walther@60723
   986
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60723
   987
\<close> ML \<open>
Walther@60723
   988
           switch_specify_solve p_ (pt, ip);
Walther@60723
   989
\<close> ML \<open>
Walther@60723
   990
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60723
   991
\<close> ML \<open>
Walther@60723
   992
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60723
   993
\<close> ML \<open>
Walther@60723
   994
           specify_do_next (pt, input_pos);
Walther@60723
   995
\<close> ML \<open>
Walther@60723
   996
"~~~~~ fun specify_do_next , args:"; val (ptp as (pt, (p, p_))) = (pt, input_pos);
Walther@60723
   997
\<close> ML \<open>
Walther@60723
   998
    val (_, (p_', tac)) =
Walther@60728
   999
   Specify.find_next_step ptp;
Walther@60723
  1000
\<close> ML \<open>
Walther@60723
  1001
"~~~~~ fun find_next_step , args:"; val (pt, pos as (_, p_)) = (ptp);
Walther@60723
  1002
\<close> ML \<open>
Walther@60723
  1003
    val {meth = met, origin = origin as (oris, o_refs as (_, pI', mI'), _), probl = pbl,
Walther@60723
  1004
      spec = refs, ...} = Calc.specify_data (pt, pos);
Walther@60723
  1005
    val ctxt = Ctree.get_ctxt pt pos;
Walther@60723
  1006
\<close> ML \<open>
Walther@60723
  1007
(*+*)val (1, [1, 2, 3], true, "#Given", Cor ((Const ("Input_Descript.Constants", _), ts ), _)) :: _
Walther@60723
  1008
  = pbl
Walther@60723
  1009
(*+*)val "[[r = 7]]" = UnparseC.terms @{context} ts;
Walther@60723
  1010
(*-----ML-^------^-HOL*)
Walther@60723
  1011
\<close> ML \<open>
Walther@60723
  1012
\<close> ML \<open>
Walther@60723
  1013
      (*if*) Ctree.just_created (pt, pos) andalso origin <> Ctree.e_origin (*else*); 
Walther@60723
  1014
\<close> ML \<open>
Walther@60723
  1015
        (*if*) p_ = Pos.Pbl (*then*); 
Walther@60723
  1016
\<close> ML \<open>
Walther@60728
  1017
           for_problem ctxt oris (o_refs, refs) (pbl, met);
Walther@60723
  1018
\<close> ML \<open>
Walther@60723
  1019
"~~~~~ fun for_problem , args:"; val (ctxt, oris, ((dI', pI', mI'), (dI, pI, mI)), (pbl, met)) =
Walther@60723
  1020
  (ctxt, oris, (o_refs, refs), (pbl, met));
Walther@60723
  1021
\<close> ML \<open>
Walther@60723
  1022
\<close> ML \<open>
Walther@60723
  1023
    val cpI = if pI = Problem.id_empty then pI' else pI;
Walther@60723
  1024
    val cmI = if mI = MethodC.id_empty then mI' else mI;
Walther@60723
  1025
    val {model = pbt, where_rls, where_, ...} = Problem.from_store ctxt cpI;
Walther@60723
  1026
    val {model = mpc, ...} = MethodC.from_store ctxt cmI
Walther@60723
  1027
\<close> ML \<open>
Walther@60723
  1028
    val (preok, _) =
Walther@60723
  1029
Pre_Conds.check_OLD ctxt where_rls where_ (pbt, I_Model.OLD_to_TEST pbl);
Walther@60723
  1030
\<close> ML \<open>
Walther@60723
  1031
"~~~~~ fun check_OLD , args:"; val (ctxt, where_rls, pres, (model_patt, i_model)) =
Walther@60723
  1032
  (ctxt, where_rls, where_, (pbt, I_Model.OLD_to_TEST pbl));
Walther@60723
  1033
\<close> ML \<open> (*ERROR lhs called with [r = 7]*)
Walther@60729
  1034
      val (_, (** )_( **)env_model(**), (env_subst, env_eval)) = of_max_variant model_patt i_model;
Walther@60729
  1035
\<close> ML \<open>
Walther@60729
  1036
(*maxi*)val "[\"\n(fixes, [[r = 7]])\"]" = env_model |> Subst.to_string @{context}
Walther@60729
  1037
(*maxi*)val "[\"\n(fixes, r)\"]" = env_subst |> Subst.to_string @{context}
Walther@60723
  1038
\<close> ML \<open>
Walther@60723
  1039
"~~~~~ fun of_max_variant , args:"; val (model_patt, i_model) = (model_patt, i_model);
Walther@60723
  1040
\<close> ML \<open>
Walther@60723
  1041
    val all_variants =
Walther@60723
  1042
        map (fn (_, variants, _, _, _) => variants) i_model
Walther@60723
  1043
        |> flat
Walther@60723
  1044
        |> distinct op =
Walther@60723
  1045
    val variants_separated = map (filter_variants' i_model) all_variants
Walther@60723
  1046
    val sums_corr = map (cnt_corrects) variants_separated
Walther@60723
  1047
    val sum_variant_s = arrange_args1 sums_corr (1, all_variants)
Walther@60723
  1048
    val (_, max_variant) = hd (*..crude decision, up to improvement *)
Walther@60723
  1049
      (sort (fn ((i, _), (j, _)) => int_ord (i, j)) sum_variant_s)
Walther@60723
  1050
    val i_model_max =
Walther@60723
  1051
      filter (fn (_, variants, _ , _ ,_) => member (op =) variants max_variant) i_model
Walther@60723
  1052
    val equal_descr_pairs = map (get_equal_descr i_model) model_patt |> flat
Walther@60723
  1053
    val env_model = make_env_model equal_descr_pairs
Walther@60723
  1054
    val equal_givens = filter (fn ((m_field, _), _) => m_field = "#Given") equal_descr_pairs
Walther@60723
  1055
\<close> ML \<open>
Walther@60723
  1056
    val subst_eval_list =
Walther@60728
  1057
           make_envs_preconds equal_givens;
Walther@60723
  1058
\<close> ML \<open>
Walther@60723
  1059
"~~~~~ fun make_envs_preconds , args:"; val (equal_givens) = (equal_givens);
Walther@60723
  1060
\<close> ML \<open>
Walther@60723
  1061
val xxx = (fn ((_, (_, id)), (_, _, _, _, (feedb, _))) =>
Walther@60723
  1062
           discern_feedback id feedb)
Walther@60728
  1063
val           ((_, (_, id)), (_, _, _, _, (feedb, _))) = nth 1 equal_givens;
Walther@60723
  1064
\<close> ML \<open>
Walther@60723
  1065
"~~~~~ fun discern_feedback , args:"; val (id, (Model_Def.Cor_TEST ((descr, ts), _))) = (id, feedb);
Walther@60723
  1066
\<close> ML \<open>
Walther@60723
  1067
           discern_typ id (descr, ts);
Walther@60723
  1068
\<close> ML \<open>
Walther@60723
  1069
"~~~~~ fun discern_typ , args:"; val (id, (descr, ts)) = (id, (descr, ts));
Walther@60723
  1070
\<close> ML \<open> (*|------------ contine me_add_find_Constants -----------------------------------------*)
Walther@60723
  1071
(*|------------------- contine me_add_find_Constants -----------------------------------------*)
Walther@60723
  1072
(*\------------------- step into me_add_find_Constants -------------------------------------//*)
Walther@60723
  1073
\<close> ML \<open> (*\------------ step into me_add_find_Constants -------------------------------------//*)
Walther@60723
  1074
\<close> ML \<open> (*while make Test_Theory work*)
Walther@60723
  1075
val (p,_,f,nxt,_,pt) = return_me_add_find_Constants;
Walther@60723
  1076
\<close> ML \<open> (*while make Test_Theory work*)
Walther@60715
  1077
(*/########################## before destroying elementwise input of lists ##################\* )
Walther@60715
  1078
val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "AdditionalValues [u]" = nxt;
Walther@60715
  1079
val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "AdditionalValues [v]" = nxt;
Walther@60715
  1080
( *\########################## before destroying elementwise input of lists ##################/*)
Walther@60715
  1081
val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "AdditionalValues [u, v]" = nxt;
Walther@60715
  1082
Walther@60715
  1083
val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Relation "Extremum (A = 2 * u * v - u \<up> 2)" = nxt;
Walther@60715
  1084
val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Relation "SideConditions [(u / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]" = nxt;
Walther@60724
  1085
Walther@60715
  1086
val return_me_Add_Relation_SideConditions
Walther@60715
  1087
                     = me nxt p c pt;
Walther@60724
  1088
(*+*)val (_, _, _, Specify_Theory "Diff_App", _, _) = return_me_Add_Relation_SideConditions; (*###############*)
Walther@60715
  1089
\<close> ML \<open>(*/------------- step into me Add_Relation_SideConditions ----------------------------\\*)
Walther@60715
  1090
(*/------------------- step into me Add_Relation_SideConditions ----------------------------\\*)
Walther@60724
  1091
\<close> ML \<open>
Walther@60715
  1092
"~~~~~ fun me , args:"; val (tac, p, _, pt) = (nxt, p, c, pt);
Walther@60724
  1093
      val ctxt = Ctree.get_ctxt pt p;
Walther@60724
  1094
\<close> ML \<open>
Walther@60724
  1095
(**)  val (pt, p) = (**) 
Walther@60724
  1096
  	    (**)case(**) Step.by_tactic tac (pt, p) (**)of(**)
Walther@60724
  1097
(**) 		    ("ok", (_, _, ptp)) => ptp (**) ;
Walther@60724
  1098
\<close> ML \<open> (*ERROR "helpless"*)
Walther@60715
  1099
    (*case*)
Walther@60715
  1100
      Step.do_next p ((pt, Pos.e_pos'), []) (*of*);
Walther@60715
  1101
\<close> ML \<open> (*//----------- step into do_next ---------------------------------------------------\\*)
Walther@60715
  1102
(*//------------------ step into do_next ---------------------------------------------------\\*)
Walther@60715
  1103
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis))
Walther@60715
  1104
  = (p, ((pt, Pos.e_pos'), [])) (*of*);
Walther@60715
  1105
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60715
  1106
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60715
  1107
      (*case*) tacis (*of*);
Walther@60715
  1108
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60715
  1109
Walther@60724
  1110
\<close> ML \<open> (*ERROR "helpless"*)
Walther@60715
  1111
      Step.switch_specify_solve p_ (pt, ip);
Walther@60724
  1112
\<close> ML \<open>
Walther@60715
  1113
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60715
  1114
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60724
  1115
\<close> ML \<open> (*ERROR "failure"*)
Walther@60715
  1116
      Step.specify_do_next (pt, input_pos);
Walther@60724
  1117
\<close> ML \<open>
Walther@60715
  1118
"~~~~~ fun specify_do_next , args:"; val ((ptp as (pt, (p, p_)))) = (pt, input_pos);
Walther@60724
  1119
\<close> ML \<open> (*ERROR "failure"*)
Walther@60724
  1120
(*isa------ERROR: Refine_Problem INSTEAD 
Walther@60724
  1121
            isa2: Specify_Theory "Diff_App"*)
Walther@60724
  1122
    val (_, (p_', tac as Specify_Theory "Diff_App")) =
Walther@60728
  1123
(*ERROR------------------^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^*)
Walther@60724
  1124
   Specify.find_next_step ptp;
Walther@60723
  1125
\<close> ML \<open>
Walther@60715
  1126
"~~~~~ fun find_next_step , args:"; val (pt, pos as (_, p_)) = (ptp);
Walther@60715
  1127
    val {meth = met, origin = origin as (oris, o_refs as (_, pI', mI'), _), probl = pbl,
Walther@60715
  1128
      spec = refs, ...} = Calc.specify_data (pt, pos);
Walther@60715
  1129
    val ctxt = Ctree.get_ctxt pt pos;
Walther@60715
  1130
      (*if*) Ctree.just_created (pt, pos) andalso origin <> Ctree.e_origin (*else*);
Walther@60724
  1131
\<close> ML \<open>
Walther@60715
  1132
        (*if*) p_ = Pos.Pbl (*then*);
Walther@60724
  1133
\<close> ML \<open>  (*ERROR Refine_Problem INSTEAD Specify_Theory "Diff_App"*)
Walther@60724
  1134
val ("dummy", (Pbl, tac as Specify_Theory "Diff_App")) =
Walther@60728
  1135
(*ERROR------------------^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^*)
Walther@60724
  1136
          for_problem ctxt oris (o_refs, refs) (pbl, met);
Walther@60715
  1137
\<close> ML \<open>
Walther@60715
  1138
"~~~~~ fun for_problem , args:"; val (ctxt, oris, ((dI', pI', mI'), (dI, pI, mI)), (pbl, met)) =
Walther@60715
  1139
  (ctxt, oris, (o_refs, refs), (pbl, met));
Walther@60715
  1140
    val cpI = if pI = Problem.id_empty then pI' else pI;
Walther@60715
  1141
    val cmI = if mI = MethodC.id_empty then mI' else mI;
Walther@60715
  1142
    val {model = pbt, where_rls, where_, ...} = Problem.from_store ctxt cpI;
Walther@60715
  1143
    val {model = mpc, ...} = MethodC.from_store ctxt cmI
Walther@60715
  1144
Walther@60715
  1145
(*+*)val [Const ("Orderings.ord_class.less", _) $ Const ("Groups.zero_class.zero", _) $
Walther@60715
  1146
  Free ("fixes", _)] = where_
Walther@60715
  1147
Walther@60724
  1148
\<close> ML \<open> (*ERROR false*)
Walther@60728
  1149
    val (preok, _) =
Walther@60724
  1150
 Pre_Conds.check_OLD ctxt where_rls where_ (pbt, I_Model.OLD_to_TEST pbl);
Walther@60723
  1151
\<close> ML \<open>
Walther@60715
  1152
\<close> ML \<open> (*///---------- step into check_OLD -------------------------------------------------\\*)
Walther@60715
  1153
(*///----------------- step into check_OLD -------------------------------------------------\\*)
Walther@60715
  1154
"~~~~~ fun check_OLD , args:"; val (ctxt, where_rls, pres, (model_patt, i_model)) =
Walther@60715
  1155
  (ctxt, where_rls, where_, (pbt, I_Model.OLD_to_TEST pbl));
Walther@60715
  1156
\<close> ML \<open>
Walther@60715
  1157
(*+*)val "[0 < fixes]" = pres |> UnparseC.terms @{context}
Walther@60715
  1158
(*+*)val "[\"(#Given, (Constants, fixes))\", \"(#Find, (Maximum, maxx))\", \"(#Find, (AdditionalValues, vals))\", \"(#Relate, (Extremum, extr))\", \"(#Relate, (SideConditions, sideconds))\"]"
Walther@60715
  1159
(*+*)  = model_patt |> Model_Pattern.to_string @{context}
Walther@60733
  1160
\<close> ML \<open>
Walther@60733
  1161
(*+*)val "[\n(1, [1, 2, 3], true ,#Given, (Cor_TEST Constants [r = 7] , pen2str, Position.T)), \n(2, [1, 2, 3], true ,#Find, (Cor_TEST Maximum A , pen2str, Position.T)), \n(3, [1, 2, 3], true ,#Find, (Cor_TEST AdditionalValues [u, v] , pen2str, Position.T)), \n(4, [1, 2, 3], true ,#Relate, (Cor_TEST Extremum (A = 2 * u * v - u \<up> 2) , pen2str, Position.T)), \n(5, [1, 2], true ,#Relate, (Cor_TEST SideConditions [(u / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2] , pen2str, Position.T))]"
Walther@60733
  1162
 = i_model |> I_Model.to_string_TEST @{context}
Walther@60723
  1163
\<close> ML \<open>
Walther@60724
  1164
val return_of_max_variant as (_, _, (env_subst, env_eval)) =
Walther@60723
  1165
           of_max_variant model_patt i_model
Walther@60715
  1166
\<close> ML \<open>
Walther@60729
  1167
(*maxi*)val "[\"\n(fixes, [[r = 7]])\"]" = env_model |> Subst.to_string @{context}
Walther@60729
  1168
(*maxi*)val "[\"\n(fixes, r)\"]" = env_subst |> Subst.to_string @{context}
Walther@60724
  1169
(*+*)val [(Free ("fixes", T1), Free ("r", T2))] = env_subst (*?!? Const*)
Walther@60724
  1170
(*+*)val Type ("Real.real", []) = T1
Walther@60724
  1171
(*+*)val Type ("Real.real", []) = T2
Walther@60715
  1172
\<close> ML \<open>
Walther@60724
  1173
(*+*)val [(Free ("r", T2), Const ("Num.numeral_class.numeral", _) $ _)] = env_eval
Walther@60724
  1174
(*+*)val Type ("Real.real", []) = T2
Walther@60715
  1175
\<close> ML \<open>
Walther@60723
  1176
\<close> ML \<open> (*while make Test_Theory work*)
Walther@60723
  1177
val (_, _, (env_subst, env_eval)) = return_of_max_variant;
Walther@60724
  1178
\<close> ML \<open> (*|||---------- contine check_OLD -----------------------------------------------------*)
Walther@60724
  1179
(*|||----------------- contine check_OLD -----------------------------------------------------*)
Walther@60715
  1180
\<close> ML \<open>
Walther@60724
  1181
      val pres_subst = map (TermC.subst_atomic_all env_subst) pres;
Walther@60715
  1182
\<close> ML \<open>
Walther@60715
  1183
\<close> ML \<open> (*|||---------- contine check_OLD -----------------------------------------------------*)
Walther@60715
  1184
(*|||----------------- contine check_OLD -----------------------------------------------------*)
Walther@60723
  1185
\<close> ML \<open>
Walther@60724
  1186
(*+*)val [(true, Const ("Orderings.ord_class.less", _) $
Walther@60724
  1187
  Const ("Groups.zero_class.zero", _) $ Free ("r", _))] = pres_subst
Walther@60724
  1188
Walther@60715
  1189
\<close> ML \<open>
Walther@60724
  1190
      val full_subst = map (fn (bool, t) => (bool, subst_atomic env_eval t)) pres_subst
Walther@60723
  1191
\<close> ML \<open>
Walther@60724
  1192
(*+*)val [(true,
Walther@60724
  1193
     Const ("Orderings.ord_class.less", _) $ Const ("Groups.zero_class.zero", _) $
Walther@60724
  1194
       (Const ("Num.numeral_class.numeral", _) $ _))] = full_subst
Walther@60715
  1195
Walther@60715
  1196
      val evals = map (eval ctxt where_rls) full_subst
Walther@60724
  1197
val return_check_OLD = (foldl and_ (true, map fst evals), pres_subst)
Walther@60724
  1198
\<close> ML \<open> (*\\\---------- step into check_OLD -------------------------------------------------\\*)
Walther@60724
  1199
(*\\\----------------- step into check_OLD -------------------------------------------------\\*)
Walther@60715
  1200
\<close> ML \<open>
Walther@60724
  1201
    val (preok as true, _) = return_check_OLD
Walther@60724
  1202
(*+---------------^^^^*)
Walther@60724
  1203
\<close> ML \<open> (*\\----------- step into do_next ---------------------------------------------------\\*)
Walther@60724
  1204
(*\\------------------ step into do_next ---------------------------------------------------\\*)
Walther@60715
  1205
\<close> ML \<open>(*\------------- step into me_Add_Relation_SideConditions ----------------------------//*)
Walther@60715
  1206
(*\------------------- step into me_Add_Relation_SideConditions ----------------------------//*)
Walther@60723
  1207
\<close> ML \<open> (*while make Test_Theory work*)
Walther@60715
  1208
val (p, _, f, nxt, _, pt) = return_me_Add_Relation_SideConditions
Walther@60715
  1209
                                      val Specify_Theory "Diff_App" = nxt;
Walther@60715
  1210
Walther@60715
  1211
val return_me_Specify_Theory
Walther@60715
  1212
                     = me nxt p c pt; val Specify_Problem ["univariate_calculus", "Optimisation"] = #4 return_me_Specify_Theory;
Walther@60715
  1213
(*/------------------- step into me Specify_Theory -----------------------------------------\\*)
Walther@60724
  1214
\<close> ML \<open> (*/------------ step into me Specify_Theory -----------------------------------------\\*)
Walther@60724
  1215
\<close> ML \<open>
Walther@60715
  1216
"~~~~~ fun me , args:"; val (tac, p, _, pt) = (nxt, p, c, pt);
Walther@60724
  1217
      val ctxt = Ctree.get_ctxt pt p;
Walther@60724
  1218
\<close> ML \<open>
Walther@60724
  1219
(*      val (pt, p) = *)
Walther@60724
  1220
  	    (*case*) Step.by_tactic tac (pt, p) (*of*);
Walther@60724
  1221
(*		    ("ok", (_, _, ptp)) => ptp *)
Walther@60724
  1222
val return_Step_by_tactic =
Walther@60728
  1223
      Step.by_tactic tac (pt, p);
Walther@60724
  1224
\<close> ML \<open> (*//----------- step into Step_by_tactic --------------------------------------------\\*)
Walther@60724
  1225
(*//------------------ step into Step_by_tactic --------------------------------------------\\*)
Walther@60724
  1226
\<close> ML \<open>
Walther@60724
  1227
"~~~~~ fun by_tactic , args:"; val (tac, (ptp as (pt, p))) = (tac, (pt, p));
Walther@60724
  1228
\<close> ML \<open>
Walther@60724
  1229
    (*case*) check tac (pt, p) (*of*);
Walther@60724
  1230
\<close> ML \<open> (*||----------- contine Step_by_tactic ------------------------------------------------*)
Walther@60724
  1231
(*||------------------ contine Step_by_tactic ------------------------------------------------*)
Walther@60724
  1232
(*\\------------------ step into Step_by_tactic --------------------------------------------//*)
Walther@60724
  1233
\<close> ML \<open> (*\\----------- step into Step_by_tactic --------------------------------------------//*)
Walther@60724
  1234
\<close> ML \<open>
Walther@60724
  1235
val ("ok", (_, _, ptp)) = return_Step_by_tactic;
Walther@60724
  1236
(*|------------------- continue me Specify_Theory --------------------------------------------*)
Walther@60724
  1237
\<close> ML \<open> (*|------------ continue me Specify_Theory --------------------------------------------*)
Walther@60724
  1238
\<close> ML \<open>
Walther@60715
  1239
val ("ok", (ts as (_, _, _) :: _, _, _)) =
Walther@60715
  1240
    (*case*)
Walther@60715
  1241
      Step.do_next p ((pt, Pos.e_pos'), []) (*of*);
Walther@60724
  1242
\<close> ML \<open> (*//----------- step into Step_do_next ----------------------------------------------\\*)
Walther@60724
  1243
(*//------------------ step into Step_do_next ----------------------------------------------\\*)
Walther@60715
  1244
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis))
Walther@60715
  1245
  = (p, ((pt, Pos.e_pos'), [])) (*of*);
Walther@60715
  1246
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60715
  1247
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60715
  1248
val _ =
Walther@60715
  1249
      (*case*) tacis (*of*);
Walther@60715
  1250
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60715
  1251
Walther@60715
  1252
      Step.switch_specify_solve p_ (pt, ip);
Walther@60715
  1253
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60715
  1254
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60715
  1255
Walther@60715
  1256
      Step.specify_do_next (pt, input_pos);
Walther@60715
  1257
"~~~~~ fun specify_do_next , args:"; val ((ptp as (pt, (p, p_)))) = (pt, input_pos);
Walther@60715
  1258
    val (_, (p_', tac)) = Specify.find_next_step ptp
Walther@60715
  1259
    val ist_ctxt =  Ctree.get_loc pt (p, p_);
Walther@60715
  1260
    (*case*) tac (*of*);
Walther@60715
  1261
Walther@60724
  1262
\<close> ML \<open>
Walther@60724
  1263
Step_Specify.by_tactic_input tac (pt, (p, p_'));
Walther@60728
  1264
(*+*)val Specify_Theory "Diff_App" = tac;
Walther@60724
  1265
\<close> ML \<open>
Walther@60724
  1266
"~~~~~ fun by_tactic_input , args:"; val ((Tactic.Specify_Theory dI), (pt, pos as (_, Pbl)))
Walther@60715
  1267
  = (tac, (pt, (p, p_')));
Walther@60715
  1268
      val (oris, dI, dI', pI', probl, ctxt) = case get_obj I pt p of
Walther@60715
  1269
        PblObj {origin = (oris, (dI,_,_),_), spec = (dI',pI',_), probl, ctxt, ...} =>
Walther@60715
  1270
          (oris, dI, dI', pI', probl, ctxt)
Walther@60715
  1271
	    val thy = ThyC.get_theory ctxt (if dI' = ThyC.id_empty then dI else dI');
Walther@60715
  1272
      val {model, where_, where_rls,...} = Problem.from_store (Ctree.get_ctxt pt pos) pI
Walther@60715
  1273
(*\\------------------ step into do_next ---------------------------------------------------//*)
Walther@60715
  1274
(*\------------------- step into me Specify_Theory -----------------------------------------//*)
Walther@60715
  1275
val (p,_,f,nxt,_,pt) = return_me_Specify_Theory;
Walther@60715
  1276
Walther@60715
  1277
val return_me_Specify_Problem (* keep for continuing me *)
Walther@60715
  1278
                     = me nxt p c pt; val Specify_Method ["Optimisation", "by_univariate_calculus"] = #4 return_me_Specify_Problem;
Walther@60715
  1279
(*/------------------- step into me Specify_Problem ----------------------------------------\\*)
Walther@60715
  1280
"~~~~~ fun me , args:"; val (tac, p, _, pt) = (nxt, p, c, pt);
Walther@60715
  1281
      val ctxt = Ctree.get_ctxt pt p
Walther@60715
  1282
Walther@60715
  1283
(** )  val ("ok", (_, _, ptp as (pt, p))) =( **)
Walther@60715
  1284
(**)    val return_by_tactic =(**) (*case*)
Walther@60715
  1285
      Step.by_tactic tac (pt, p) (*of*);
Walther@60715
  1286
(*//------------------ step into by_tactic -------------------------------------------------\\*)
Walther@60715
  1287
"~~~~~ fun by_tactic , args:"; val (tac, (ptp as (pt, p))) = (tac, (pt, p));
Walther@60715
  1288
Walther@60715
  1289
    (*case*)
Walther@60715
  1290
      Step.check tac (pt, p) (*of*);
Walther@60715
  1291
"~~~~~ fun check , args:"; val (tac, (ctree, pos)) = (tac, (pt, p));
Walther@60715
  1292
  (*if*) Tactic.for_specify tac (*then*);
Walther@60715
  1293
Walther@60715
  1294
Specify_Step.check tac (ctree, pos);
Walther@60715
  1295
"~~~~~ fun check , args:"; val ((Tactic.Specify_Problem pID), (pt, pos as (p, _)))
Walther@60715
  1296
  = (tac, (ctree, pos));
Walther@60715
  1297
		    val (oris, dI, pI, dI', pI', itms) = case Ctree.get_obj I pt p of
Walther@60715
  1298
		      Ctree.PblObj {origin = (oris, (dI, pI, _), _), spec = (dI', pI', _), probl = itms, ...}
Walther@60715
  1299
		        => (oris, dI, pI, dI', pI', itms)
Walther@60715
  1300
	      val thy = ThyC.get_theory ctxt (if dI' = ThyC.id_empty then dI else dI');
Walther@60715
  1301
(*\\------------------ step into by_tactic -------------------------------------------------//*)
Walther@60715
  1302
val ("ok", (_, _, ptp as (pt, p))) = return_by_tactic (* kept for continuing me *);
Walther@60715
  1303
Walther@60715
  1304
      (*case*)
Walther@60715
  1305
      Step.do_next p ((pt, Pos.e_pos'), []) (*of*);
Walther@60715
  1306
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis)) = (p, ((pt, Pos.e_pos'), []));
Walther@60715
  1307
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60715
  1308
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60715
  1309
val _ =
Walther@60715
  1310
      (*case*) tacis (*of*);
Walther@60715
  1311
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60715
  1312
Walther@60715
  1313
      Step.switch_specify_solve p_ (pt, ip);
Walther@60715
  1314
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60715
  1315
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60715
  1316
Walther@60715
  1317
      Step.specify_do_next (pt, input_pos);
Walther@60715
  1318
"~~~~~ fun specify_do_next , args:"; val ((ptp as (pt, (p, p_)))) = ((pt, input_pos));
Walther@60715
  1319
    val (_, (p_', tac)) = Specify.find_next_step ptp
Walther@60715
  1320
    val ist_ctxt =  Ctree.get_loc pt (p, p_)
Walther@60715
  1321
val _ =
Walther@60715
  1322
    (*case*) tac (*of*);
Walther@60715
  1323
Walther@60715
  1324
Step_Specify.by_tactic_input tac (pt, (p, p_'));
Walther@60715
  1325
"~~~~~ fun by_tactic_input , args:"; val ((Tactic.Specify_Method id), (pt, pos))
Walther@60715
  1326
  = (tac, (pt, (p, p_')));
Walther@60715
  1327
Walther@60715
  1328
      val (o_model, ctxt, i_model) =
Walther@60715
  1329
Specify_Step.complete_for id (pt, pos);
Walther@60715
  1330
"~~~~~ fun complete_for , args:"; val (mID, (ctree, pos)) = (id, (pt, pos));
Walther@60715
  1331
    val {origin = (o_model, o_refs, _), spec = refs, probl = i_prob, ctxt,
Walther@60715
  1332
       ...} = Calc.specify_data (ctree, pos);
Walther@60715
  1333
    val ctxt = Ctree.get_ctxt ctree pos
Walther@60715
  1334
    val (dI, _, _) = References.select_input o_refs refs;
Walther@60715
  1335
    val {model = m_patt, where_, where_rls, ...} = MethodC.from_store ctxt mID
Walther@60715
  1336
    val {origin = (root_model, _, _), ...} = Calc.specify_data (ctree, ([], Pos.Und))
Walther@60715
  1337
    val (o_model', ctxt') = O_Model.complete_for m_patt root_model (o_model, ctxt)
Walther@60715
  1338
    val thy = ThyC.get_theory ctxt dI
Walther@60715
  1339
    val (_, (i_model, _)) = M_Match.match_itms_oris ctxt i_prob (m_patt, where_, where_rls) o_model';
Walther@60715
  1340
(*\------------------- step into me Specify_Problem ----------------------------------------//*)
Walther@60715
  1341
val (p,_,f,nxt,_,pt) = return_me_Specify_Problem
Walther@60715
  1342
Walther@60715
  1343
val return_me_Add_Given_FunctionVariable
Walther@60715
  1344
                     = me nxt p c pt; val Add_Given "FunctionVariable a" = #4 return_me_Add_Given_FunctionVariable;
Walther@60715
  1345
(*/------------------- step into me Specify_Method -----------------------------------------\\*)
Walther@60715
  1346
"~~~~~ fun me , args:"; val (tac, p, _, pt) = (nxt, p, c, pt);
Walther@60715
  1347
      val ctxt = Ctree.get_ctxt pt p
Walther@60715
  1348
      val (pt, p) = 
Walther@60715
  1349
  	    (*Step.by_tactic is here for testing by me; do_next would suffice in me*)
Walther@60715
  1350
  	    case Step.by_tactic tac (pt, p) of
Walther@60715
  1351
  		    ("ok", (_, _, ptp)) => ptp;
Walther@60715
  1352
Walther@60715
  1353
         (*case*)
Walther@60715
  1354
      Step.do_next p ((pt, Pos.e_pos'), []) (*of*);
Walther@60715
  1355
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis)) = (p, ((pt, Pos.e_pos'), []));
Walther@60715
  1356
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60715
  1357
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60715
  1358
val _ =
Walther@60715
  1359
      (*case*) tacis (*of*);
Walther@60715
  1360
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60715
  1361
Walther@60715
  1362
      Step.switch_specify_solve p_ (pt, ip);
Walther@60715
  1363
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60715
  1364
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60715
  1365
Walther@60715
  1366
      Step.specify_do_next (pt, input_pos);
Walther@60715
  1367
"~~~~~ fun specify_do_next , args:"; val ((ptp as (pt, (p, p_)))) = (pt, input_pos);
Walther@60715
  1368
Walther@60715
  1369
    val (_, (p_', tac)) =
Walther@60715
  1370
   Specify.find_next_step ptp;
Walther@60715
  1371
"~~~~~ fun find_next_step , args:"; val ((pt, pos as (_, p_))) = (ptp);
Walther@60715
  1372
    val {meth = met, origin = origin as (oris, o_refs as (_, pI', mI'), _), probl = pbl,
Walther@60715
  1373
      spec = refs, ...} = Calc.specify_data (pt, pos);
Walther@60715
  1374
    val ctxt = Ctree.get_ctxt pt pos;
Walther@60715
  1375
      (*if*) Ctree.just_created (pt, pos) andalso origin <> Ctree.e_origin (*else*);
Walther@60715
  1376
        (*if*) p_ = Pos.Pbl (*else*);
Walther@60715
  1377
Walther@60730
  1378
\<close> ML \<open>
Walther@60715
  1379
   Specify.for_method ctxt oris (o_refs, refs) (pbl, met);
Walther@60730
  1380
\<close> ML \<open>
Walther@60730
  1381
"~~~~~ fun for_method , args:"; val (ctxt, oris, ((dI', pI', mI'), (dI, pI, mI)), (_, met))
Walther@60715
  1382
  = (ctxt, oris, (o_refs, refs), (pbl, met));
Walther@60730
  1383
\<close> ML \<open>
Walther@60715
  1384
    val cmI = if mI = MethodC.id_empty then mI' else mI;
Walther@60715
  1385
    val {model = mpc, where_rls, where_, ...} = MethodC.from_store ctxt cmI;    (*..MethodC ?*)
Walther@60730
  1386
    val (preok, _) = Pre_Conds.check_OLD ctxt where_rls where_ (mpc, I_Model.OLD_to_TEST met);
Walther@60715
  1387
val NONE =
Walther@60715
  1388
    (*case*) find_first (I_Model.is_error o #5) met (*of*);
Walther@60715
  1389
Walther@60730
  1390
\<close> ML \<open>
Walther@60715
  1391
      (*case*)
Walther@60715
  1392
   Specify.item_to_add (ThyC.get_theory ctxt 
Walther@60715
  1393
     (if dI = ThyC.id_empty then dI' else dI)) oris mpc met (*of*);
Walther@60730
  1394
\<close> ML \<open>
Walther@60715
  1395
"~~~~~ fun item_to_add , args:"; val (thy, oris, _, itms)
Walther@60715
  1396
  = ((ThyC.get_theory ctxt (if dI = ThyC.id_empty then dI' else dI)), oris, mpc, met);
Walther@60730
  1397
\<close> ML \<open>
Walther@60715
  1398
(*\------------------- step into me Specify_Method -----------------------------------------//*)
Walther@60715
  1399
val (p,_,f,nxt,_,pt) = return_me_Add_Given_FunctionVariable
Walther@60715
  1400
Walther@60715
  1401
val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "Input_Descript.Domain {0<..<r}" = nxt;
Walther@60715
  1402
val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "ErrorBound (\<epsilon> = 0)" = nxt;
Walther@60724
  1403
\<close> ML \<open> (*/------------ step into me_Add_Given_ErrorBound------------------------------------\\*)
Walther@60724
  1404
(* ------------------- step into UNTIL A PROGRAM IS REQUIRED (not yet impl. ---------------- )*)
Walther@60724
  1405
(*/------------------- step into me_Add_Given_ErrorBound------------------------------------\\*)
Walther@60724
  1406
\<close> ML \<open>
Walther@60724
  1407
"~~~~~ fun me , args:"; val (tac, p, _(*NEW remove*), pt) = (nxt, p, c, pt);
Walther@60724
  1408
\<close> ML \<open>
Walther@60724
  1409
      val ctxt = Ctree.get_ctxt pt p
Walther@60724
  1410
      val (pt, p) = 
Walther@60724
  1411
  	    (*Step.by_tactic is here for testing by me; do_next would suffice in me*)
Walther@60724
  1412
  	    case Step.by_tactic tac (pt, p) of
Walther@60728
  1413
  		    ("ok", (_, _, ptp)) => ptp;
Walther@60724
  1414
\<close> ML \<open>
Walther@60728
  1415
(*ERROR due to missing program in creating the environment from formal args* )
Walther@60724
  1416
      (*case*)
Walther@60724
  1417
      Step.do_next p ((pt, Pos.e_pos'), []) (*of*);
Walther@60724
  1418
( *ERROR*)
Walther@60724
  1419
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis) ) =
Walther@60724
  1420
  (p, ((pt, Pos.e_pos'), []));
Walther@60724
  1421
\<close> ML \<open>
Walther@60724
  1422
  (*if*) Pos.on_calc_end ip (*else*); 
Walther@60724
  1423
\<close> ML \<open>
Walther@60724
  1424
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60724
  1425
\<close> ML \<open>
Walther@60724
  1426
val _ =
Walther@60724
  1427
      (*case*) tacis (*of*);
Walther@60724
  1428
\<close> ML \<open>
Walther@60724
  1429
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60724
  1430
\<close> ML \<open>
Walther@60728
  1431
(*ERROR due to missing program in creating the environment from formal args* )
Walther@60728
  1432
           switch_specify_solve p_ (pt, ip);
Walther@60724
  1433
( *ERROR*)
Walther@60728
  1434
\<close> ML \<open>
Walther@60724
  1435
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60724
  1436
\<close> ML \<open>
Walther@60724
  1437
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60724
  1438
\<close> ML \<open>
Walther@60728
  1439
(*ERROR due to missing program in creating the environment from formal args* )
Walther@60724
  1440
           specify_do_next (pt, input_pos)
Walther@60724
  1441
( *ERROR*)
Walther@60724
  1442
"~~~~~ fun specify_do_next , args:"; val ((ptp as (pt, (p, p_)))) = ((pt, input_pos));
Walther@60724
  1443
\<close> ML \<open>
Walther@60724
  1444
    val (_, (p_', tac as Apply_Method ["Optimisation", "by_univariate_calculus"])) =
Walther@60724
  1445
    (*-------------------^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ not yet impl.*)
Walther@60724
  1446
      Specify.find_next_step ptp
Walther@60724
  1447
    val ist_ctxt =  Ctree.get_loc pt (p, p_)
Walther@60724
  1448
\<close> ML \<open>
Walther@60724
  1449
val Tactic.Apply_Method (mI as ["Optimisation", "by_univariate_calculus"]) =
Walther@60724
  1450
        (*case*) tac (*of*);
Walther@60724
  1451
\<close> ML \<open>
Walther@60728
  1452
(*ERROR due to missing program in creating the environment from formal args* )
Walther@60724
  1453
  	    LI.by_tactic (Tactic.Apply_Method' (mI, NONE, Istate.empty, ContextC.empty))
Walther@60724
  1454
  	      ist_ctxt (pt, (p, p_'))
Walther@60724
  1455
( *ERROR*)
Walther@60724
  1456
"~~~~~ fun by_tactic , args:"; val ((Tactic.Apply_Method' (mI, _, _, _)), (_, ctxt), (pt, pos as (p, _))) =
Walther@60724
  1457
  ((Tactic.Apply_Method' (mI, NONE, Istate.empty, ContextC.empty)), ist_ctxt, (pt, (p, p_')));
Walther@60724
  1458
\<close> ML \<open>
Walther@60724
  1459
      val (itms, oris, probl) = case get_obj I pt p of
Walther@60724
  1460
          PblObj {meth = itms, origin = (oris, _, _), probl, ...} => (itms, oris, probl)
Walther@60724
  1461
      val {model, ...} = MethodC.from_store ctxt mI;
Walther@60724
  1462
\<close> ML \<open>
Walther@60724
  1463
      (*if*) itms <> [] (*then*);
Walther@60724
  1464
\<close> ML \<open>
Walther@60724
  1465
      val itms = I_Model.complete oris probl [] model
Walther@60724
  1466
\<close> ML \<open>
Walther@60724
  1467
(*+*)val "[\n(1 ,[1, 2, 3] ,true ,#Given ,Cor Constants [r = 7] ,(fixes, [[r = 7]])), \n(2 ,[1, 2, 3] ,true ,#Find ,Cor Maximum A ,(maxx, [A])), \n(3 ,[1, 2, 3] ,true ,#Find ,Cor AdditionalValues [u, v] ,(vals, [[u, v]])), \n(4 ,[1, 2, 3] ,true ,#Relate ,Cor Extremum (A = 2 * u * v - u \<up> 2) ,(extr, [A = 2 * u * v - u \<up> 2])), \n(5 ,[1, 2] ,true ,#Relate ,Cor SideConditions [(u / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2] ,(sideconds, [[(u / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]])), \n(7 ,[1] ,true ,#Given ,Cor FunctionVariable a ,(funvar, [a])), \n(10 ,[1, 2] ,true ,#Given ,Cor Input_Descript.Domain {0<..<r} ,(doma, [{0<..<r}])), \n(12 ,[1, 2, 3] ,true ,#Given ,Cor ErrorBound (\<epsilon> = 0) ,(err, [\<epsilon> = 0]))]"
Walther@60724
  1468
 = itms |> I_Model.to_string @{context}
Walther@60724
  1469
(*+*)val 8 = length itms
Walther@60724
  1470
(*+*)val 8 = length model
Walther@60724
  1471
(*\------------------- step into me_Add_Given_ErrorBound------------------------------------//*)
Walther@60724
  1472
\<close> ML \<open> (*\------------ step into me_Add_Given_ErrorBound------------------------------------//*)
Walther@60710
  1473
\<close> ML \<open>
Walther@60710
  1474
\<close> ML \<open>
Walther@60710
  1475
\<close>
Walther@60710
  1476
Walther@60732
  1477
section \<open>=====equa "Minisubpbl/100-init-rootpbl-NEXT_STEP.sml" =================================\<close>
Walther@60726
  1478
ML \<open>
Walther@60726
  1479
\<close> ML \<open>
Walther@60726
  1480
(* Title:  "Minisubpbl/100-init-rootpbl-NEXT_STEP.sml"
Walther@60726
  1481
   Author: Walther Neuper 2307
Walther@60726
  1482
*)
Walther@60729
  1483
(*overwrites NEW funs in Test_Theory ABOVE* )
Walther@60726
  1484
open Ctree;
Walther@60726
  1485
open Pos;
Walther@60726
  1486
open TermC;
Walther@60726
  1487
open Istate;
Walther@60726
  1488
open Tactic;
Walther@60726
  1489
open I_Model;
Walther@60726
  1490
open Pre_Conds;
Walther@60726
  1491
open Rewrite;
Walther@60726
  1492
open Step;
Walther@60726
  1493
open Specify;
Walther@60726
  1494
open LItool;
Walther@60726
  1495
open LI;
Walther@60729
  1496
( *overwrites NEW funs in Test_Theory ABOVE*)
Walther@60726
  1497
Walther@60726
  1498
"----------- Minisubpbl/100-init-rootpbl-NEXT_STEP.sml ---------------------------------------";
Walther@60726
  1499
"----------- Minisubpbl/100-init-rootpbl-NEXT_STEP.sml ---------------------------------------";
Walther@60726
  1500
"----------- Minisubpbl/100-init-rootpbl-NEXT_STEP.sml ---------------------------------------";
Walther@60726
  1501
tracing "--- Minisubplb/100-init-rootpbl-NEXT_STEP.sml ---------------------------------------";
Walther@60726
  1502
val fmz = ["equality (x+1=(2::real))", "solveFor x", "solutions L"];
Walther@60728
  1503
val (dI',pI',mI') =
Walther@60726
  1504
  ("Test", ["sqroot-test", "univariate", "equation", "test"],
Walther@60726
  1505
   ["Test", "squ-equ-test-subpbl1"]);
Walther@60726
  1506
val (p,_,f,nxt,_,pt) = Test_Code.init_calc @{context} [(fmz, (dI',pI',mI'))];
Walther@60726
  1507
(*[], Pbl*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, e_pos'), []); val Model_Problem = tac
Walther@60726
  1508
\<close> ML \<open>
Walther@60726
  1509
val return_do_next_Model_Problem as xxx = 
Walther@60726
  1510
      Step.do_next p ((pt, e_pos'), []);
Walther@60726
  1511
\<close> ML \<open>
Walther@60726
  1512
val return_do_next_Model_Problem as ("ok", ([(Specify_Theory "Test", _, _)], _, _))
Walther@60726
  1513
  = Step.do_next p ((pt, e_pos'), []);
Walther@60726
  1514
\<close> ML \<open> (*/------------ step into do_next_Model_Problem -------------------------------------\\*)
Walther@60726
  1515
(*/------------------- step into do_next_Model_Problem -------------------------------------\\*)
Walther@60726
  1516
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis)) = (p, ((pt, e_pos'), []));
Walther@60726
  1517
\<close> ML \<open>
Walther@60726
  1518
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60726
  1519
\<close> ML \<open>
Walther@60726
  1520
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60726
  1521
\<close> ML \<open>
Walther@60726
  1522
val _ =
Walther@60726
  1523
      (*case*) tacis (*of*);
Walther@60726
  1524
\<close> ML \<open>
Walther@60726
  1525
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60726
  1526
\<close> ML \<open>
Walther@60726
  1527
           switch_specify_solve p_ (pt, ip);
Walther@60726
  1528
\<close> ML \<open>
Walther@60726
  1529
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60726
  1530
\<close> ML \<open>
Walther@60726
  1531
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60726
  1532
\<close> ML \<open>
Walther@60726
  1533
           specify_do_next (pt, input_pos);
Walther@60726
  1534
\<close> ML \<open>
Walther@60726
  1535
"~~~~~ fun specify_do_next , args:"; val ((ptp as (pt, (p, p_)))) = ((pt, input_pos));
Walther@60726
  1536
\<close> ML \<open>
Walther@60726
  1537
    val (_, (p_', tac as Specify_Theory "Test")) =
Walther@60726
  1538
   Specify.find_next_step ptp;
Walther@60726
  1539
\<close> ML \<open>
Walther@60726
  1540
"~~~~~ fun find_next_step , args:"; val ((pt, pos as (_, p_))) = (ptp);
Walther@60726
  1541
\<close> ML \<open>
Walther@60726
  1542
    val {meth = met, origin = origin as (oris, o_refs as (_, pI', mI'), _), probl = pbl,
Walther@60726
  1543
      spec = refs, ...} = Calc.specify_data (pt, pos);
Walther@60726
  1544
    val ctxt = Ctree.get_ctxt pt pos
Walther@60726
  1545
\<close> ML \<open>
Walther@60726
  1546
      (*if*) Ctree.just_created (pt, pos) andalso origin <> Ctree.e_origin (*else*);
Walther@60726
  1547
\<close> ML \<open>
Walther@60726
  1548
        (*if*) p_ = Pos.Pbl (*then*);
Walther@60726
  1549
\<close> ML \<open>
Walther@60726
  1550
           for_problem ctxt oris (o_refs, refs) (pbl, met);
Walther@60726
  1551
\<close> ML \<open>
Walther@60726
  1552
"~~~~~ fun for_problem , args:"; val (ctxt, oris, ((dI', pI', mI'), (dI, pI, mI)), (pbl, met)) =
Walther@60726
  1553
  (ctxt, oris, (o_refs, refs), (pbl, met));
Walther@60726
  1554
\<close> ML \<open>
Walther@60726
  1555
    val cpI = if pI = Problem.id_empty then pI' else pI;
Walther@60726
  1556
    val cmI = if mI = MethodC.id_empty then mI' else mI;
Walther@60726
  1557
    val {model = pbt, where_rls, where_, ...} = Problem.from_store ctxt cpI;
Walther@60726
  1558
    val {model = mpc, ...} = MethodC.from_store ctxt cmI
Walther@60726
  1559
\<close> ML \<open> (*ERROR?---vvvvv*)
Walther@60726
  1560
    val (preok as true, xxx) =
Walther@60726
  1561
 Pre_Conds.check_OLD ctxt where_rls where_ (pbt, I_Model.OLD_to_TEST pbl);
Walther@60726
  1562
\<close> ML \<open>
Walther@60733
  1563
(*+*)val "[\n(1, [1], true ,#Given, (Cor_TEST equality (x + 1 = 2) , pen2str, Position.T)), \n(2, [1], true ,#Given, (Cor_TEST solveFor x , pen2str, Position.T)), \n(3, [1], true ,#Find, (Cor_TEST solutions L , pen2str, Position.T))]"
Walther@60733
  1564
 = I_Model.OLD_to_TEST pbl |> I_Model.to_string_TEST @{context};
Walther@60726
  1565
(*+*)val "[precond_rootpbl v_v]" = where_ |> UnparseC.terms @{context};
Walther@60726
  1566
\<close> ML \<open>
Walther@60726
  1567
"~~~~~ fun check_OLD , args:"; val (ctxt, where_rls, pres, (model_patt, i_model)) =
Walther@60726
  1568
  (ctxt, where_rls, where_, (pbt, I_Model.OLD_to_TEST pbl));
Walther@60726
  1569
\<close> ML \<open>
Walther@60726
  1570
      val (_, env_model, (env_subst, env_eval)) = of_max_variant model_patt i_model
Walther@60726
  1571
\<close> ML \<open>
Walther@60729
  1572
(*equa*)val "[\"\n(e_e, x + 1 = 2)\", \"\n(v_v, x)\", \"\n(v_v'i', L)\"]" = env_model |> Subst.to_string @{context}
Walther@60729
  1573
(*equa*)val "[]" = env_subst |> Subst.to_string @{context}
Walther@60729
  1574
\<close> ML \<open>
Walther@60726
  1575
      val pres_subst = map (TermC.subst_atomic_all env_subst) pres;
Walther@60726
  1576
\<close> ML \<open>
Walther@60726
  1577
(*+*)val [(false, Const ("Test.precond_rootpbl", _) $ Free ("v_v", _))] = pres_subst
Walther@60726
  1578
\<close> ML \<open>
Walther@60726
  1579
      val pres_subst_other = map (TermC.subst_atomic_all env_model) (map #2 pres_subst);
Walther@60726
  1580
\<close> ML \<open>
Walther@60726
  1581
(*+*)val "[\"\n(e_e, x + 1 = 2)\", \"\n(v_v, x)\", \"\n(v_v'i', L)\"]"
Walther@60726
  1582
  = env_model |> Subst.to_string @{context}
Walther@60726
  1583
(*+*)val [(true, Const ("Test.precond_rootpbl", _) $ Free ("x", _))] = pres_subst_other
Walther@60726
  1584
\<close> ML \<open>
Walther@60726
  1585
      val full_subst = if env_eval = [] then pres_subst_other
Walther@60726
  1586
        else map (TermC.subst_atomic_all env_eval) (map #2 pres_subst_other)
Walther@60726
  1587
\<close> ML \<open>
Walther@60726
  1588
(*+*)val [] = env_eval
Walther@60726
  1589
(*+*)val [(true, Const ("Test.precond_rootpbl", _) $ Free ("x", _))] = full_subst
Walther@60726
  1590
\<close> ML \<open>
Walther@60726
  1591
      val evals = map (eval ctxt where_rls) full_subst
Walther@60726
  1592
\<close> ML \<open>
Walther@60726
  1593
(*+*)val [(true, Const ("Test.precond_rootpbl", _) $ Free ("x", _))] = evals
Walther@60726
  1594
\<close> ML \<open>
Walther@60726
  1595
val return_check_OLD = (foldl and_ (true, map fst evals), pres_subst_other)
Walther@60726
  1596
\<close> ML \<open>
Walther@60726
  1597
(*+*)val (true, [(true, Const ("Test.precond_rootpbl", _) $ Free ("x", _))]) = return_check_OLD
Walther@60726
  1598
\<close> ML \<open>
Walther@60726
  1599
\<close> ML \<open>
Walther@60726
  1600
\<close> ML \<open> (*|------------ contine do_next_Model_Problem -----------------------------------------*)
Walther@60726
  1601
(*|------------------- contine do_next_Model_Problem -----------------------------------------*)
Walther@60726
  1602
(*\------------------- step into do_next_Model_Problem -------------------------------------//*)
Walther@60726
  1603
\<close> ML \<open> (*\------------ step into do_next_Model_Problem -------------------------------------//*)
Walther@60726
  1604
val (_, ([(tac as Specify_Theory "Test", _, _)], _, (pt, p))) = return_do_next_Model_Problem;
Walther@60726
  1605
\<close> ML \<open>
Walther@60726
  1606
(*[], Pbl*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, e_pos'), []); val Specify_Problem ["sqroot-test", "univariate", "equation", "test"] = tac
Walther@60726
  1607
Walther@60726
  1608
(*[], Met*)val return_Step_do_next_Specify_Problem = Step.do_next p ((pt, e_pos'), []);(*Specify_Method ["Test", "squ-equ-test-subpbl1"]*)
Walther@60726
  1609
(*/------------------- step into Step.do_next_Specify_Problem ------------------------------\\*)
Walther@60726
  1610
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis)) =
Walther@60726
  1611
  (p, ((pt, e_pos'), []));
Walther@60726
  1612
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60726
  1613
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60726
  1614
val _ =
Walther@60726
  1615
      (*case*) tacis (*of*);
Walther@60726
  1616
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60726
  1617
Walther@60726
  1618
\<close> ML \<open>
Walther@60726
  1619
      Step.switch_specify_solve p_ (pt, ip);
Walther@60726
  1620
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60726
  1621
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60726
  1622
Walther@60726
  1623
      Step.specify_do_next (pt, input_pos);
Walther@60726
  1624
"~~~~~ fun specify_do_next , args:"; val ((ptp as (pt, (p, p_)))) = ((pt, input_pos));
Walther@60726
  1625
    val (_, (p_', tac)) = Specify.find_next_step ptp
Walther@60726
  1626
    val ist_ctxt =  Ctree.get_loc pt (p, p_)
Walther@60726
  1627
val _ =
Walther@60726
  1628
    (*case*) tac (*of*);
Walther@60726
  1629
Walther@60726
  1630
Step_Specify.by_tactic_input tac (pt, (p, p_'));
Walther@60726
  1631
"~~~~~ fun by_tactic_input , args:"; val ((Tactic.Specify_Method id), (pt, pos)) =
Walther@60726
  1632
  (tac, (pt, (p, p_')));
Walther@60726
  1633
Walther@60726
  1634
\<close> ML \<open>
Walther@60726
  1635
      val (o_model, ctxt, i_model) =
Walther@60726
  1636
Specify_Step.complete_for id (pt, pos);
Walther@60726
  1637
"~~~~~ fun complete_for , args:"; val (mID, (ctree, pos)) = (id, (pt, pos));
Walther@60726
  1638
    val {origin = (o_model, _, _), probl = i_prob, ctxt,
Walther@60726
  1639
       ...} = Calc.specify_data (ctree, pos);
Walther@60726
  1640
    val {model = m_patt, where_, where_rls, ...} = MethodC.from_store ctxt mID
Walther@60726
  1641
    val {origin = (root_model, _, _), ...} = Calc.specify_data (ctree, ([], Pos.Und))
Walther@60726
  1642
    val (o_model', ctxt') = O_Model.complete_for m_patt root_model (o_model, ctxt)
Walther@60726
  1643
Walther@60726
  1644
    val (_, (i_model, _)) =
Walther@60726
  1645
   M_Match.match_itms_oris ctxt i_prob (m_patt, where_, where_rls) o_model';
Walther@60726
  1646
"~~~~~ fun match_itms_oris , args:"; val (ctxt, pbl, (pbt, where_, where_rls), oris) =
Walther@60726
  1647
  (ctxt, i_prob, (m_patt, where_, where_rls), o_model');
Walther@60726
  1648
 (*0*)val mv = Pre_Conds.max_variant pbl;
Walther@60726
  1649
Walther@60726
  1650
      fun eqdsc_pbt_itm ((_,(d,_))) (_, _, _, _, itm_) = d = I_Model.descriptor itm_;
Walther@60726
  1651
      fun notmem pbl pbt1 = case find_first (eqdsc_pbt_itm pbt1) pbl of
Walther@60726
  1652
				SOME _ => false | NONE => true;
Walther@60726
  1653
 (*1*)val mis = (filter (notmem pbl)) pbt;
Walther@60726
  1654
      fun eqdsc_ori (_,(d,_)) (_, _, _, d', _) = d = d';
Walther@60726
  1655
      fun ori2itmMis (f, (d, pid)) (i, v, _, _, _) = (i, v, false, f, I_Model.Mis (d, pid));
Walther@60726
  1656
 (*2*)fun oris2itms oris mis1 = ((map (ori2itmMis mis1)) o (filter (eqdsc_ori mis1))) oris;
Walther@60726
  1657
      val news = (flat o (map (oris2itms oris))) mis;
Walther@60726
  1658
 (*3*)fun mem_vat (_, vats, _, _, _) = member op = vats mv;
Walther@60726
  1659
      val newitms = filter mem_vat news;
Walther@60726
  1660
 (*4*)val itms' = pbl @ newitms;
Walther@60726
  1661
Walther@60726
  1662
\<close> ML \<open>
Walther@60726
  1663
      val (pb, where_') =
Walther@60726
  1664
 Pre_Conds.check_OLD ctxt where_rls where_ (m_patt, I_Model.OLD_to_TEST itms');
Walther@60726
  1665
\<close> ML \<open>
Walther@60726
  1666
"~~~~~ fun check_OLD , args:"; val (ctxt, where_rls, pres, (model_patt, i_model)) =
Walther@60726
  1667
  (ctxt, where_rls, where_, (m_patt, I_Model.OLD_to_TEST itms'));
Walther@60726
  1668
      val (_, env_model, (env_subst, env_eval)) = of_max_variant model_patt i_model
Walther@60729
  1669
\<close> ML \<open>
Walther@60729
  1670
(*equa*)val "[\"\n(e_e, x + 1 = 2)\", \"\n(v_v, x)\", \"\n(v_v'i', L)\"]" = env_model |> Subst.to_string @{context}
Walther@60729
  1671
(*equa*)val "[]" = env_subst |> Subst.to_string @{context}
Walther@60729
  1672
\<close> ML \<open>
Walther@60726
  1673
      val pres_subst = map (TermC.subst_atomic_all env_subst) pres;
Walther@60726
  1674
      val pres_subst_other = map (TermC.subst_atomic_all env_model) (map #2 pres_subst);
Walther@60726
  1675
      val full_subst = if env_eval = [] then pres_subst_other
Walther@60726
  1676
        else map (TermC.subst_atomic_all env_eval) (map #2 pres_subst_other)
Walther@60726
  1677
Walther@60726
  1678
\<close> ML \<open>
Walther@60726
  1679
(*+*)val "Test" = ctxt |> Proof_Context.theory_of  |> ThyC.id_of
Walther@60726
  1680
(*+*)val Rule_Set.Repeat {id = "prls_met_test_squ_sub", ...} = where_rls
Walther@60726
  1681
(*+*)val [(true, tTEST as Const ("Test.precond_rootmet", _) $ Free ("x", xxx))] = pres_subst_other
Walther@60726
  1682
Walther@60726
  1683
\<close> ML \<open>
Walther@60726
  1684
\<close> ML \<open>
Walther@60726
  1685
(*+*)val ctxt = Config.put rewrite_trace true ctxt;
Walther@60726
  1686
      val evals = map (
Walther@60726
  1687
 Pre_Conds.eval ctxt where_rls) full_subst;
Walther@60726
  1688
(*
Walther@60726
  1689
@## rls: prls_met_test_squ_sub on: precond_rootmet x 
Walther@60726
  1690
@### try calc: "Test.precond_rootmet" 
Walther@60726
  1691
@#### eval asms: "(precond_rootmet x) = True" 
Walther@60726
  1692
@### calc. to: True 
Walther@60726
  1693
@### try calc: "Test.precond_rootmet" 
Walther@60726
  1694
@### try calc: "Test.precond_rootmet" 
Walther@60726
  1695
*)
Walther@60726
  1696
(*+*)val ctxt = Config.put rewrite_trace false ctxt;
Walther@60726
  1697
Walther@60726
  1698
\<close> ML \<open>
Walther@60726
  1699
"~~~~~ fun eval_precond_rootmet , args:"; val ((thmid:string), _, (t as (Const (\<^const_name>\<open>Test.precond_rootmet\<close>, _) $ arg)), ctxt) =
Walther@60726
  1700
  ("aaa", "bbb", tTEST, ctxt);
Walther@60726
  1701
    SOME (TermC.mk_thmid thmid (UnparseC.term ctxt arg) "",
Walther@60726
  1702
      HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
Walther@60726
  1703
;
Walther@60726
  1704
    (TermC.mk_thmid thmid (UnparseC.term ctxt arg)) "";
Walther@60726
  1705
\<close> ML \<open>
Walther@60726
  1706
"~~~~~ fun mk_thmid , args:"; val (thmid, n1, n2) = (thmid, (UnparseC.term ctxt arg), "");
Walther@60726
  1707
  thmid ^ (cut_longid n1) ^ "_" ^ (cut_longid n2);
Walther@60726
  1708
Walther@60726
  1709
\<close> ML \<open>
Walther@60726
  1710
      (cut_longid n2);
Walther@60726
  1711
"~~~~~ fun cut_longid , args:"; val (dn) = (n2);
Walther@60726
  1712
(*\------------------- step into Step.do_next_Specify_Problem ------------------------------//*)
Walther@60726
  1713
val (_, ([(tac, _, _)], _, (pt, p))) = return_Step_do_next_Specify_Problem; val Specify_Method ["Test", "squ-equ-test-subpbl1"] = tac;
Walther@60726
  1714
Walther@60727
  1715
\<close> ML \<open>
Walther@60726
  1716
(*[1], Frm*)val return_Step_do_next_Specify_Method = Step.do_next p ((pt, e_pos'), []);(*Apply_Method ["Test", "squ-equ-test-subpbl1"]*)
Walther@60726
  1717
\<close> ML \<open>(*/------------- step into do_next_Specify_Method ------------------------------------\\*)
Walther@60726
  1718
(*/------------------- step into do_next_Specify_Method ------------------------------------\\*)
Walther@60726
  1719
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis)) =
Walther@60726
  1720
  (p, ((pt, e_pos'), []));
Walther@60726
  1721
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60726
  1722
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60726
  1723
val _ =
Walther@60726
  1724
      (*case*) tacis (*of*);
Walther@60726
  1725
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60726
  1726
Walther@60729
  1727
\<close> ML \<open> (*ERROR 2 formal args: [e_e, v_v], 0 actual args: []*)
Walther@60726
  1728
           switch_specify_solve p_ (pt, ip);
Walther@60726
  1729
\<close> ML \<open>
Walther@60726
  1730
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60726
  1731
      (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60726
  1732
Walther@60729
  1733
\<close> ML \<open> (*ERROR 2 formal args: [e_e, v_v], 0 actual args: []*)
Walther@60726
  1734
           specify_do_next (pt, input_pos);
Walther@60726
  1735
\<close> ML \<open>
Walther@60726
  1736
"~~~~~ fun specify_do_next , args:"; val (ptp as (pt, (p, p_))) = (pt, input_pos);
Walther@60726
  1737
    val (_, (p_', tac)) = Specify.find_next_step ptp
Walther@60726
  1738
    val ist_ctxt =  Ctree.get_loc pt (p, p_)
Walther@60726
  1739
(*+*)val Tactic.Apply_Method mI =
Walther@60726
  1740
    (*case*) tac (*of*);
Walther@60726
  1741
Walther@60729
  1742
\<close> ML \<open> (*ERROR 2 formal args: [e_e, v_v], 0 actual args: []*)
Walther@60726
  1743
  	    LI.by_tactic (Tactic.Apply_Method' (mI, NONE, Istate.empty, ContextC.empty))
Walther@60726
  1744
  	      ist_ctxt (pt, (p, p_'));
Walther@60726
  1745
\<close> ML \<open>
Walther@60726
  1746
"~~~~~ fun by_tactic , args:"; val ((Tactic.Apply_Method' (mI, _, _, _)), (_, ctxt), (pt, pos as (p, _))) =
Walther@60726
  1747
  ((Tactic.Apply_Method' (mI, NONE, Istate.empty, ContextC.empty)),
Walther@60726
  1748
  	      ist_ctxt, (pt, (p, p_')));
Walther@60726
  1749
         val (itms, oris, probl) = case get_obj I pt p of (*STILL asms=[]*)
Walther@60726
  1750
           PblObj {meth = itms, origin = (oris, _, _), probl, ...} => (itms, oris, probl)
Walther@60726
  1751
         | _ => raise ERROR "LI.by_tactic Apply_Method': uncovered case get_obj"
Walther@60726
  1752
         val {model, ...} = MethodC.from_store ctxt mI;
Walther@60726
  1753
         val itms = if itms <> [] then itms else I_Model.complete oris probl [] model
Walther@60726
  1754
Walther@60727
  1755
\<close> ML \<open>
Walther@60727
  1756
         val return_init_pstate = case
Walther@60726
  1757
           LItool.init_pstate ctxt (I_Model.OLD_to_TEST itms) mI of
Walther@60726
  1758
             (is as Pstate {env, ...}, ctxt, program) => (is, env, ctxt, program)
Walther@60726
  1759
(*#------------------- un-hide local of init_pstate -----------------------------------------\*)
Walther@60726
  1760
fun msg_miss ctxt sc metID caller f formals actuals =
Walther@60726
  1761
  "ERROR in creating the environment from formal args. of partial_function \"" ^ fst (Program.get_fun_id sc) ^ "\"\n" ^
Walther@60726
  1762
	"and the actual args., ie. items of the guard of \"" ^ MethodC.id_to_string metID ^ "\" by \"" ^ caller ^ "\":\n" ^
Walther@60726
  1763
	"formal arg \"" ^ UnparseC.term ctxt f ^ "\" doesn't match any actual arg.\n" ^
Walther@60726
  1764
	"with:\n" ^
Walther@60726
  1765
	(string_of_int o length) formals ^ " formal args: " ^ UnparseC.terms ctxt formals ^ "\n" ^
Walther@60726
  1766
	(string_of_int o length) actuals ^ " actual args: " ^ UnparseC.terms ctxt actuals
Walther@60726
  1767
fun msg_miss_type ctxt sc metID f formals actuals =
Walther@60726
  1768
  "ERROR in creating the environment from formal args. of partial_function \"" ^ fst (Program.get_fun_id sc) ^ "\"\n" ^
Walther@60726
  1769
	"and the actual args., ie. items of the guard of \"" ^ MethodC.id_to_string metID ^ "\" by \"assoc_by_type\":\n" ^
Walther@60726
  1770
	"formal arg \"" ^ UnparseC.term ctxt f ^ "\" of type \"" ^ UnparseC.typ ctxt (type_of f) ^
Walther@60726
  1771
  "\" doesn't mach any actual arg.\nwith:\n" ^
Walther@60726
  1772
	(string_of_int o length) formals ^ " formal args: " ^ UnparseC.terms ctxt formals ^ "\n" ^
Walther@60726
  1773
	(string_of_int o length) actuals ^ " actual args: " ^ UnparseC.terms ctxt actuals ^ "\n" ^
Walther@60726
  1774
  "   with types: " ^ (strs2str o (map ((UnparseC.typ ctxt) o type_of))) actuals
Walther@60726
  1775
fun msg_ambiguous ctxt sc metID f aas formals actuals =
Walther@60726
  1776
  "AMBIGUITY in creating the environment from formal args. of partial_function \"" ^ fst (Program.get_fun_id sc) ^ "\"\n" ^
Walther@60726
  1777
	"and the actual args., ie. items of the guard of \"" ^ MethodC.id_to_string metID ^ "\" by \"assoc_by_type\":\n" ^
Walther@60726
  1778
  "formal arg. \"" ^ UnparseC.term ctxt f ^ "\" type-matches with several..."  ^
Walther@60726
  1779
  "actual args. \"" ^ UnparseC.terms ctxt aas ^ "\"\n" ^
Walther@60726
  1780
  "selected \"" ^ UnparseC.term ctxt (hd aas) ^ "\"\n" ^
Walther@60726
  1781
	"with\n" ^
Walther@60726
  1782
	"formals: " ^ UnparseC.terms ctxt formals ^ "\n" ^
Walther@60726
  1783
	"actuals: " ^ UnparseC.terms ctxt actuals
Walther@60726
  1784
fun trace_init ctxt metID =
Walther@60726
  1785
  if Config.get ctxt LI_trace
Walther@60726
  1786
  then tracing("@@@ program \"" ^ strs2str metID ^ "\"")
Walther@60726
  1787
  else ();
Walther@60726
  1788
Walther@60726
  1789
fun assoc_by_type ctxt f aa prog met_id formals actuals =
Walther@60726
  1790
  case filter (curry (fn (f, a) => type_of f = type_of a) f) aa of
Walther@60726
  1791
    [] => raise ERROR (msg_miss_type ctxt prog met_id f formals actuals)
Walther@60726
  1792
  | [a] => (f, a)
Walther@60726
  1793
  | aas as (a :: _) => (writeln (msg_ambiguous ctxt prog met_id f aas formals actuals); (f, a));
Walther@60726
  1794
(*
Walther@60726
  1795
  fun assoc_formals_actuals: at "PIDE turn 11" ONLY syntax revised, NOT semantics
Walther@60726
  1796
  Env.T ->         (* [] for building return Env.T         *)
Walther@60726
  1797
  term list ->     (* changed during building return Env.T *)
Walther@60726
  1798
  term list ->     (* changed during building return Env.T *)
Walther@60726
  1799
  Proof.context -> (*                                      *)
Walther@60726
  1800
  term ->          (* program code                         *)
Walther@60726
  1801
  MethodC.id ->    (* for error msg                        *)
Walther@60726
  1802
  term list ->     (* original value, unchanged            *)
Walther@60726
  1803
  term list ->     (* original value, unchanged            *)
Walther@60726
  1804
  Env.T            (* return Env.T                         *)
Walther@60726
  1805
*)
Walther@60726
  1806
\<close> ML \<open>
Walther@60726
  1807
fun relate_args _ (f::_) [] ctxt prog met_id formals actuals =
Walther@60726
  1808
    raise ERROR (msg_miss ctxt prog met_id "relate_args" f formals actuals)
Walther@60726
  1809
  | relate_args env [] _ _ _ _ _ _ = env (*may drop Find?*)
Walther@60726
  1810
  | relate_args env (f::ff) (aas as (a::aa)) ctxt prog met_id formals actuals = 
Walther@60726
  1811
    if type_of f = type_of a 
Walther@60726
  1812
    then relate_args (env @ [(f, a)]) ff aa ctxt prog met_id formals actuals
Walther@60726
  1813
    else
Walther@60726
  1814
      let val (f, a) = assoc_by_type ctxt f aas prog met_id formals actuals
Walther@60726
  1815
      in relate_args (env @ [(f, a)]) ff (remove op = a aas) ctxt prog met_id formals actuals end
Walther@60726
  1816
;
Walther@60726
  1817
(*+*)relate_args: Env.T -> term list -> term list -> Proof.context -> term -> MethodC.id ->
Walther@60726
  1818
    term list -> term list -> Env.T;
Walther@60726
  1819
(*#------------------- un-hide local of init_pstate -----------------------------------------/*)
Walther@60726
  1820
Walther@60727
  1821
\<close> ML \<open>
Walther@60727
  1822
\<close> ML \<open>
Walther@60726
  1823
\<close> ML \<open>(*//------------ step into init_pstate -----------------------------------------------\\*)
Walther@60726
  1824
(*//------------------ step into init_pstate -----------------------------------------------\\*)
Walther@60726
  1825
val (_, ctxt, i_model, met_id) = (prog_rls, ctxt, itms, mI);
Walther@60726
  1826
"~~~~~ fun init_pstate , args:"; val (ctxt, i_model, met_id) =
Walther@60726
  1827
    (ctxt, I_Model.OLD_to_TEST i_model, met_id);
Walther@60726
  1828
    val (model_patt, program, prog, prog_rls, where_, where_rls) =
Walther@60726
  1829
      case MethodC.from_store ctxt met_id of
Walther@60726
  1830
        {model = model_patt, program = program as Rule.Prog prog, prog_rls, where_, where_rls,...} =>
Walther@60726
  1831
          (model_patt, program, prog, prog_rls, where_, where_rls)
Walther@60726
  1832
      | _ => raise ERROR ("init_pstate with " ^ MethodC.id_to_string met_id)
Walther@60726
  1833
Walther@60726
  1834
\<close> ML \<open>
Walther@60726
  1835
    val return_of_max_variant = 
Walther@60729
  1836
           of_max_variant model_patt i_model;
Walther@60726
  1837
\<close> ML \<open>(*///----------- step into of_max_variant --------------------------------------------\\*)
Walther@60726
  1838
(*///----------------- step into of_max_variant --------------------------------------------\\*)
Walther@60726
  1839
"~~~~~ fun of_max_variant , args:"; val (model_patt, i_model) = (model_patt, i_model);
Walther@60726
  1840
Walther@60726
  1841
(*+*)i_model: I_Model.T_TEST;
Walther@60726
  1842
Walther@60726
  1843
    val all_variants =
Walther@60726
  1844
        map (fn (_, variants, _, _, _) => variants) i_model
Walther@60726
  1845
        |> flat
Walther@60726
  1846
        |> distinct op =
Walther@60726
  1847
    val variants_separated = map (filter_variants' i_model) all_variants
Walther@60726
  1848
    val sums_corr = map (cnt_corrects) variants_separated
Walther@60726
  1849
(*+*)val [3] = sums_corr;
Walther@60726
  1850
Walther@60726
  1851
    val sum_variant_s = arrange_args1 sums_corr (1, all_variants)
Walther@60726
  1852
(*+*)val [(3, 1)] = sum_variant_s;
Walther@60726
  1853
Walther@60726
  1854
    val (_, max_variant) = hd (*..crude decision, up to improvement *)
Walther@60726
  1855
      (sort (fn ((i, _), (j, _)) => int_ord (i, j)) sum_variant_s)
Walther@60726
  1856
(*+*)val 1 = max_variant;
Walther@60726
  1857
Walther@60726
  1858
    val i_model_max =
Walther@60726
  1859
      filter (fn (_, variants, _ , _ ,_) => member (op =) variants max_variant) i_model
Walther@60726
  1860
(*+*)val 3 = length i_model_max;
Walther@60726
  1861
Walther@60726
  1862
    val equal_descr_pairs =
Walther@60726
  1863
      map (get_equal_descr i_model) model_patt
Walther@60726
  1864
      |> flat
Walther@60726
  1865
(*+*)val 3 = length equal_descr_pairs; (*only 1 variant in this test*)
Walther@60726
  1866
Walther@60726
  1867
\<close> ML \<open>
Walther@60726
  1868
    val env_model = make_env_model equal_descr_pairs
Walther@60726
  1869
    val equal_givens = filter (fn ((m_field, _), _) => m_field = "#Given") equal_descr_pairs
Walther@60726
  1870
\<close> ML \<open>
Walther@60726
  1871
(*+*)val ([], []) = (env_subst, env_eval)
Walther@60726
  1872
\<close> ML \<open>
Walther@60727
  1873
(*+*)if (equal_descr_pairs |> descr_pairs_to_string @{context}) = "[" ^
Walther@60726
  1874
  "((#Given, (equality, e_e)), " ^
Walther@60733
  1875
    "(1, [1], true ,#Given, (Cor_TEST equality (x + 1 = 2) , pen2str, Position.T))), " ^
Walther@60726
  1876
  "((#Given, (solveFor, v_v)), " ^
Walther@60733
  1877
    "(2, [1], true ,#Given, (Cor_TEST solveFor x , pen2str, Position.T))), " ^
Walther@60726
  1878
  "((#Find, (solutions, v_v'i')), " ^
Walther@60733
  1879
    "(3, [1], true ,#Find, (Cor_TEST solutions L , pen2str, Position.T)))]"
Walther@60727
  1880
  then () else error "equal_descr_pairs CHANGED";
Walther@60726
  1881
\<close> ML \<open>(*|||----------- contine of_max_variant ------------------------------------------------*)
Walther@60726
  1882
(*|||----------------- contine of_max_variant ------------------------------------------------*)
Walther@60726
  1883
    val subst_eval_list = make_envs_preconds equal_givens
Walther@60726
  1884
    val (env_subst, env_eval) = split_list subst_eval_list
Walther@60726
  1885
Walther@60726
  1886
\<close> ML \<open>
Walther@60726
  1887
val return_of_max_variant_step = (i_model_max, env_model, (env_subst, env_eval));
Walther@60727
  1888
\<close> ML \<open>
Walther@60727
  1889
(*+*) val "[\"\n(e_e, x + 1 = 2)\", \"\n(v_v, x)\", \"\n(v_v'i', L)\"]"
Walther@60727
  1890
  = env_model |>Subst.to_string @{context};
Walther@60726
  1891
(*+*)if return_of_max_variant_step = return_of_max_variant then () else error "<>";
Walther@60726
  1892
\<close> ML \<open>(*\\\----------- step into of_max_variant --------------------------------------------\\*)
Walther@60726
  1893
(*\\\----------------- step into of_max_variant --------------------------------------------\\*)
Walther@60727
  1894
val (_, env_model, (env_subst as [], env_eval as [])) = return_of_max_variant
Walther@60726
  1895
 
Walther@60726
  1896
\<close> ML \<open>(*||------------ continue init_pstate ------------------------------------------------\\*)
Walther@60726
  1897
(*||------------------ continue init_pstate ------------------------------------------------\\*)
Walther@60726
  1898
    val actuals = map snd env_model
Walther@60726
  1899
\<close> ML \<open>
Walther@60726
  1900
(*+*)val "[x + 1 = 2, x, L]" = actuals |> UnparseC.terms @{context}
Walther@60726
  1901
(*+*)val [_, Free ("x", _), Free ("L", typ)] = actuals
Walther@60726
  1902
\<close> ML \<open>
Walther@60726
  1903
(*+*)val Type ("List.list", [Type ("HOL.bool", [])]) = typ
Walther@60726
  1904
\<close> ML \<open>
Walther@60726
  1905
    val formals = Program.formal_args prog
Walther@60726
  1906
(*+*)val "[e_e, v_v]" = (formals |> UnparseC.terms @{context})
Walther@60726
  1907
Walther@60726
  1908
(*+*)val "minisubpbl e_e v_v =\n(let e_ea =\n       (Try (Rewrite_Set ''norm_equation'') #>\n        Try (Rewrite_Set ''Test_simplify''))\n        e_e;\n     L_La =\n       SubProblem\n        (''Test'', [''LINEAR'', ''univariate'', ''equation'', ''test''],\n         [''Test'', ''solve_linear''])\n        [BOOL e_e, REAL v_v]\n in Check_elementwise L_L {v_v. Assumptions})" =
Walther@60726
  1909
  (prog |> UnparseC.term @{context})
Walther@60726
  1910
Walther@60726
  1911
\<close> ML \<open>
Walther@60727
  1912
    val preconds = Pre_Conds.check_envs_TEST ctxt where_rls where_ (env_model, (env_subst, env_eval))
Walther@60726
  1913
Walther@60726
  1914
\<close> ML \<open>(*||------------ continue init_pstate ------------------------------------------------\\*)
Walther@60726
  1915
(*||------------------ continue init_pstate ------------------------------------------------\\*)
Walther@60726
  1916
    val ctxt = ctxt |> ContextC.insert_assumptions (map snd (snd preconds));
Walther@60726
  1917
    val ist = Istate.e_pstate
Walther@60726
  1918
      |> Istate.set_eval prog_rls
Walther@60726
  1919
      |> Istate.set_env_true (relate_args [] formals actuals ctxt prog met_id formals actuals);
Walther@60726
  1920
(*+*)                        (relate_args [] formals actuals ctxt prog met_id formals actuals);
Walther@60726
  1921
(*+*)val "[\"\n(e_e, x + 1 = 2)\", \"\n(v_v, x)\"]" =
Walther@60726
  1922
                             (relate_args [] formals actuals ctxt prog met_id formals actuals)
Walther@60726
  1923
       |> Subst.to_string @{context}
Walther@60726
  1924
Walther@60726
  1925
val return_init_pstate_steps = (* = return_init_pstate*)
Walther@60726
  1926
    (Istate.Pstate ist, ctxt, program)
Walther@60726
  1927
\<close> ML \<open>(*\\------------ step into init_pstate -----------------------------------------------//*)
Walther@60726
  1928
(*\\------------------ step into init_pstate -----------------------------------------------//*)
Walther@60726
  1929
val (is, env, ctxt, prog) = return_init_pstate;
Walther@60726
  1930
Walther@60726
  1931
\<close> ML \<open>(*|------------- continuing do_next_Specify_Method -------------------------------------*)
Walther@60726
  1932
(*|------------------- continuing do_next_Specify_Method -------------------------------------*)
Walther@60726
  1933
\<close> ML \<open>(*\------------- step into do_next_Specify_Method ------------------------------------//*)
Walther@60726
  1934
(*\------------------- step into do_next_Specify_Method ------------------------------------//*)
Walther@60726
  1935
\<close> ML \<open>
Walther@60726
  1936
val (_, ([(tac, _, _)], _, (pt, p))) = return_Step_do_next_Specify_Method; val Apply_Method ["Test", "squ-equ-test-subpbl1"] = tac;
Walther@60726
  1937
Walther@60726
  1938
(*[1], Frm*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, e_pos'), []); val Rewrite_Set "norm_equation" = tac
Walther@60726
  1939
Walther@60726
  1940
(*[1], Res*)val return_Step_do_next_Rewrite_Set = Step.do_next p ((pt, e_pos'), []);
Walther@60726
  1941
\<close> ML \<open>(*/------------- step into do_next_Apply_Method --------------------------------------\\*)
Walther@60726
  1942
(*/------------------- step into do_next_Apply_Method --------------------------------------\\*)
Walther@60726
  1943
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis)) =
Walther@60726
  1944
  (p ,((pt, e_pos'), []));
Walther@60726
  1945
\<close> ML \<open>
Walther@60726
  1946
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60726
  1947
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60726
  1948
val _ =
Walther@60726
  1949
      (*case*) tacis (*of*);
Walther@60726
  1950
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60726
  1951
Walther@60726
  1952
      Step.switch_specify_solve p_ (pt, ip);
Walther@60726
  1953
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos))
Walther@60726
  1954
  = (p_, (pt, ip));
Walther@60726
  1955
      (*if*) Pos.on_specification ([], state_pos) (*else*);
Walther@60726
  1956
Walther@60726
  1957
        LI.do_next (pt, input_pos);
Walther@60726
  1958
"~~~~~ and do_next , args:"; val (ptp as (pt, pos as (p, p_))) = (pt, input_pos);
Walther@60726
  1959
    (*if*) MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p) (*else*);
Walther@60726
  1960
	      val ((ist, ctxt), sc) = LItool.resume_prog (p,p_) pt;
Walther@60726
  1961
Walther@60726
  1962
val return_LI_find_next_step = (*case*)
Walther@60726
  1963
        LI.find_next_step sc (pt, pos) ist ctxt (*of*);
Walther@60726
  1964
(*//------------------ step into LI.find_next_step -----------------------------------------\\*)
Walther@60726
  1965
"~~~~~ fun find_next_step , args:"; val ((Rule.Prog prog), (ptp as (pt, pos as (p, _))), (Istate.Pstate ist), ctxt) =
Walther@60726
  1966
  (sc, (pt, pos), ist, ctxt);
Walther@60726
  1967
(*.. this showed that we have ContextC.empty*)
Walther@60726
  1968
(*\\------------------ step into LI.find_next_step -----------------------------------------//*)
Walther@60726
  1969
val LI.Next_Step (ist, ctxt, tac) = return_LI_find_next_step;
Walther@60726
  1970
Walther@60726
  1971
        LI.by_tactic tac (ist, Tactic.insert_assumptions tac ctxt) ptp;
Walther@60726
  1972
"~~~~~ fun by_tactic , args:"; val (tac_, ic, (pt, pos))
Walther@60726
  1973
  = (tac, (ist, Tactic.insert_assumptions tac ctxt), ptp);
Walther@60726
  1974
      val pos = next_in_prog' pos
Walther@60726
  1975
Walther@60726
  1976
 	    val (pos', c, _, pt) =
Walther@60726
  1977
Solve_Step.add_general tac_ ic (pt, pos);
Walther@60726
  1978
"~~~~~ fun add_general , args:"; val (tac, ic, cs)
Walther@60726
  1979
  = (tac_, ic, (pt, pos));
Walther@60726
  1980
 (*if*) Tactic.for_specify' tac (*else*);
Walther@60726
  1981
Walther@60726
  1982
Solve_Step.add tac ic cs;
Walther@60726
  1983
"~~~~~ fun add , args:"; val ((Tactic.Rewrite_Set' (_, _, rls', f, (f', asm))), (is, ctxt), (pt, (p, _)))
Walther@60726
  1984
  = (tac, ic, cs);
Walther@60726
  1985
      val (pt, c) = Ctree.cappend_atomic pt p (is, ctxt) f 
Walther@60726
  1986
        (Tactic.Rewrite_Set (Rule_Set.id rls')) (f', asm) Ctree.Complete
Walther@60726
  1987
      val pt = Ctree.update_branch pt p Ctree.TransitiveB
Walther@60726
  1988
Walther@60726
  1989
val return =
Walther@60726
  1990
      ((p, Pos.Res), c, Test_Out.FormKF (UnparseC.term ctxt f'), pt)
Walther@60726
  1991
\<close> ML \<open>(*\------------- step into do_next_Apply_Method --------------------------------------//*)
Walther@60726
  1992
(*\------------------- step into do_next_Apply_Method --------------------------------------//*)
Walther@60726
  1993
val (_, ([(tac, _, _)], _, (pt, p))) = return_Step_do_next_Rewrite_Set
Walther@60726
  1994
Walther@60726
  1995
(* final test ... ----------------------------------------------------------------------------*)
Walther@60726
  1996
(*+*)val ([2], Res) = p;
Walther@60726
  1997
Test_Tool.show_pt_tac pt; (*[
Walther@60726
  1998
([], Frm), solve (x + 1 = 2, x)
Walther@60726
  1999
. . . . . . . . . . Apply_Method ["Test", "squ-equ-test-subpbl1"], 
Walther@60726
  2000
([1], Frm), x + 1 = 2
Walther@60726
  2001
. . . . . . . . . . Rewrite_Set "norm_equation", 
Walther@60726
  2002
([1], Res), x + 1 + - 1 * 2 = 0
Walther@60726
  2003
. . . . . . . . . . Rewrite_Set "Test_simplify", 
Walther@60726
  2004
([2], Res), - 1 + x = 0
Walther@60726
  2005
. . . . . . . . . . Check_Postcond ["sqroot-test", "univariate", "equation", "test"]]   ..INCORRECT
Walther@60726
  2006
*)
Walther@60727
  2007
Walther@60726
  2008
\<close> ML \<open>
Walther@60726
  2009
\<close>
Walther@60726
  2010
Walther@60726
  2011
Walther@60732
  2012
section \<open>=====bieg Interpret/li-tool.sml =======================================================\<close>
Walther@60729
  2013
ML \<open>
Walther@60729
  2014
\<close> ML \<open>
Walther@60729
  2015
(* Title:  Interpret/li-tool.sml
Walther@60729
  2016
   Author: Walther Neuper 050908
Walther@60729
  2017
   (c) copyright due to lincense terms.
Walther@60729
  2018
*)
Walther@60729
  2019
"-----------------------------------------------------------------";
Walther@60729
  2020
"table of contents -----------------------------------------------";
Walther@60729
  2021
"-----------------------------------------------------------------";
Walther@60729
  2022
"----------- fun implicit_take, fun get_first_argument -------------------------";
Walther@60729
  2023
"----------- fun from_prog ---------------------------------------";
Walther@60729
  2024
"----------- fun specific_from_prog ----------------------------";
Walther@60729
  2025
"----------- fun de_esc_underscore -------------------------------";
Walther@60729
  2026
"----------- fun go ----------------------------------------------";
Walther@60729
  2027
"----------- fun formal_args -------------------------------------";
Walther@60729
  2028
"----------- fun dsc_valT ----------------------------------------";
Walther@60729
  2029
"----------- fun arguments_from_model ---------------------------------------";
Walther@60729
  2030
"----------- fun init_pstate -----------------------------------------------------------------";
Walther@60729
  2031
"-----------------------------------------------------------------";
Walther@60729
  2032
"-----------------------------------------------------------------";
Walther@60729
  2033
"-----------------------------------------------------------------";
Walther@60729
  2034
Walther@60729
  2035
val thy =  @{theory Isac_Knowledge};
Walther@60729
  2036
Walther@60729
  2037
\<close> ML \<open>
Walther@60729
  2038
"----------- fun init_pstate -----------------------------------------------------------------";
Walther@60729
  2039
"----------- fun init_pstate -----------------------------------------------------------------";
Walther@60729
  2040
"----------- fun init_pstate -----------------------------------------------------------------";
Walther@60729
  2041
(*
Walther@60729
  2042
  This test is deeply nested (down into details of creating environments).
Walther@60729
  2043
  In order to make Sidekick show the structure add to each
Walther@60729
  2044
  *                    (*/...\*) and        (*\.../*)
Walther@60729
  2045
  * a companion  > ML < (*/...\*) and > ML < (*\.../*)
Walther@60729
  2046
  Note the wrong ^----^ delimiters.
Walther@60729
  2047
*)
Walther@60729
  2048
val fmz = ["Traegerlaenge L", "Streckenlast q_0", "Biegelinie y",
Walther@60729
  2049
	     "Randbedingungen [y 0 = (0::real), y L = 0, M_b 0 = 0, M_b L = 0]",
Walther@60729
  2050
	     "FunktionsVariable x", "GleichungsVariablen [c, c_2, c_3, c_4]",
Walther@60729
  2051
	     "AbleitungBiegelinie dy"];
Walther@60729
  2052
val (dI',pI',mI') = ("Biegelinie", ["Biegelinien"],
Walther@60729
  2053
		     ["IntegrierenUndKonstanteBestimmen2"]);
Walther@60729
  2054
val p = e_pos'; val c = [];
Walther@60729
  2055
(*[], Pbl*)val (p,_,f,nxt,_,pt) = Test_Code.init_calc @{context} [(fmz, (dI',pI',mI'))]; val Model_Problem = nxt
Walther@60729
  2056
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "Traegerlaenge L" = nxt
Walther@60729
  2057
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "Streckenlast q_0" = nxt
Walther@60729
  2058
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "Biegelinie y" = nxt
Walther@60729
  2059
\<close> ML \<open>
Walther@60729
  2060
(*/---broken elementwise input to lists---\* )
Walther@60729
  2061
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; (*\<rightarrow>Add_Relation "Randbedingungen [y 0 = 0]", ERROR MISSING step: M_b 0 = 0*)
Walther@60729
  2062
                                        (*isa* ) val Add_Relation "Randbedingungen [y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0]" = nxt
Walther@60729
  2063
                                        ( *isa2*) val Add_Relation "Randbedingungen [y 0 = 0]" = nxt
Walther@60729
  2064
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; (*\<rightarrow>Add_Relation "Randbedingungen [y L = 0, M_b 0 = 0, M_b L = 0]"*)
Walther@60729
  2065
                                        (*isa*) val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2066
                                        (*isa2* ) val Add_Relation "Randbedingungen [y L = 0, M_b 0 = 0, M_b L = 0]" = nxt( **)
Walther@60729
  2067
( *\---broken elementwise input to lists---/*)
Walther@60729
  2068
\<close> ML \<open>
Walther@60729
  2069
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Relation "Randbedingungen [y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0]" = nxt
Walther@60729
  2070
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2071
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["Biegelinien"] = nxt
Walther@60729
  2072
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["IntegrierenUndKonstanteBestimmen2"] = nxt
Walther@60729
  2073
(*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "FunktionsVariable x" = nxt
Walther@60729
  2074
(*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "GleichungsVariablen [c, c_2, c_3, c_4]" = nxt
Walther@60729
  2075
(*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "AbleitungBiegelinie dy" = nxt
Walther@60729
  2076
Walther@60729
  2077
\<close> ML \<open>
Walther@60729
  2078
(*[], Met*)val return_Add_Given_AbleitungBieg
Walther@60729
  2079
                                = me nxt p c pt;
Walther@60729
  2080
\<close> ML \<open> (*/------------ step into me_Add_Given_AbleitungBieg --------------------------------\\*)
Walther@60729
  2081
(*/------------------- step into me_Add_Given_AbleitungBieg --------------------------------\\*)
Walther@60729
  2082
Walther@60729
  2083
"~~~~~ fun me , args:"; val (tac, p, _(*NEW remove*), pt) = (nxt, p, c, pt);
Walther@60729
  2084
      val ctxt = Ctree.get_ctxt pt p
Walther@60729
  2085
val ("ok", (_, _, ptp as (pt, p))) =
Walther@60729
  2086
  	    (*case*) Step.by_tactic tac (pt, p) (*of*);
Walther@60729
  2087
Walther@60729
  2088
\<close> ML \<open>
Walther@60729
  2089
        (*case*)
Walther@60729
  2090
      Step.do_next p ((pt, Pos.e_pos'), []) (*of*);
Walther@60729
  2091
\<close> ML \<open>
Walther@60729
  2092
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis)) =
Walther@60729
  2093
  (p, ((pt, Pos.e_pos'), []) );
Walther@60729
  2094
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60729
  2095
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60729
  2096
val _ =
Walther@60729
  2097
      (*case*) tacis (*of*);
Walther@60729
  2098
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60729
  2099
Walther@60729
  2100
\<close> ML \<open>
Walther@60729
  2101
           switch_specify_solve p_ (pt, ip);
Walther@60729
  2102
\<close> ML \<open>
Walther@60729
  2103
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60729
  2104
       (*if*) Pos.on_specification ([], state_pos) (*then*);
Walther@60729
  2105
Walther@60729
  2106
\<close> ML \<open>
Walther@60729
  2107
           specify_do_next (pt, input_pos);
Walther@60729
  2108
\<close> ML \<open>
Walther@60729
  2109
"~~~~~ fun specify_do_next , args:"; val (ptp as (pt, (p, p_))) = (pt, input_pos);
Walther@60729
  2110
    val (_, (p_', tac)) = Specify.find_next_step ptp
Walther@60729
  2111
    val ist_ctxt =  Ctree.get_loc pt (p, p_)
Walther@60729
  2112
val Tactic.Apply_Method mI =
Walther@60729
  2113
    (*case*) tac (*of*);
Walther@60729
  2114
Walther@60729
  2115
\<close> ML \<open>
Walther@60729
  2116
  	    LI.by_tactic (Tactic.Apply_Method' (mI, NONE, Istate.empty, ContextC.empty))
Walther@60729
  2117
  	      ist_ctxt (pt, (p, p_'));
Walther@60729
  2118
\<close> ML \<open>
Walther@60729
  2119
"~~~~~ fun by_tactic , args:"; val ((Tactic.Apply_Method' (mI, _, _, _)), (_, ctxt), (pt, pos as (p, _))) =
Walther@60729
  2120
  ((Tactic.Apply_Method' (mI, NONE, Istate.empty, ContextC.empty)), ist_ctxt, (pt, (p, p_')));
Walther@60729
  2121
      val (itms, oris, probl) = case get_obj I pt p of (*STILL asms=[]*)
Walther@60729
  2122
          PblObj {meth = itms, origin = (oris, _, _), probl, ...} => (itms, oris, probl)
Walther@60729
  2123
        | _ => raise ERROR "LI.by_tactic Apply_Method': uncovered case get_obj"
Walther@60729
  2124
      val {model, ...} = MethodC.from_store ctxt mI;
Walther@60729
  2125
      val itms = if itms <> [] then itms else I_Model.complete oris probl [] model
Walther@60729
  2126
;
Walther@60729
  2127
(*+*)if (itms |> I_Model.to_string @{context}) = "[\n" ^
Walther@60729
  2128
  "(1 ,[1] ,true ,#Given ,Cor Traegerlaenge L ,(l_l, [L])), \n" ^
Walther@60729
  2129
  "(2 ,[1] ,true ,#Given ,Cor Streckenlast q_0 ,(q_q, [q_0])), \n" ^
Walther@60729
  2130
  "(3 ,[1] ,true ,#Find ,Cor Biegelinie y ,(b_b, [y])), \n" ^
Walther@60729
  2131
  "(4 ,[1] ,true ,#Relate ,Cor Randbedingungen [y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0] ,(r_b, [[y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0]])), \n" ^
Walther@60729
  2132
  "(5 ,[1] ,true ,#Given ,Cor FunktionsVariable x ,(fun_var, [x])), \n" ^
Walther@60729
  2133
  "(6 ,[1] ,true ,#Given ,Cor GleichungsVariablen [c, c_2, c_3, c_4] ,(vs, [[c, c_2, c_3, c_4]])), \n" ^
Walther@60729
  2134
  "(7 ,[1] ,true ,#Given ,Cor AbleitungBiegelinie dy ,(id_der, [dy]))]"
Walther@60729
  2135
(*+*)then () else error "init_pstate: initial i_model changed";
Walther@60729
  2136
(*+*)if (itms |> I_Model.penv_to_string @{context}) = "[" ^ (*to be eliminated*)
Walther@60729
  2137
(*1*)"(l_l, [L]), " ^
Walther@60729
  2138
(*2*)"(q_q, [q_0]), " ^
Walther@60729
  2139
(*3*)"(b_b, [y]), " ^
Walther@60729
  2140
(*4*)"(r_b, [[y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0]]), " ^
Walther@60729
  2141
(*5*)"(fun_var, [x]), " ^
Walther@60729
  2142
(*6*)"(vs, [[c, c_2, c_3, c_4]]), " ^
Walther@60729
  2143
(*7*)"(id_der, [dy])]"
Walther@60729
  2144
(*+*)then () else error "init_pstate: initial penv changed";
Walther@60729
  2145
Walther@60729
  2146
\<close> ML \<open>
Walther@60729
  2147
    (*case*)
Walther@60729
  2148
    LItool.init_pstate ctxt (I_Model.OLD_to_TEST itms) mI (*of*);
Walther@60729
  2149
\<close> ML \<open>(*//------------ step into init_pstate -----------------------------------------------\\*)
Walther@60729
  2150
(*//------------------ step into init_pstate -----------------------------------------------\\*)
Walther@60729
  2151
"~~~~~ fun init_pstate , args:"; val (ctxt, i_model, met_id) = (ctxt, (I_Model.OLD_to_TEST itms), mI);
Walther@60729
  2152
    val (model_patt, program, prog, prog_rls, where_, where_rls) =
Walther@60729
  2153
      case MethodC.from_store ctxt met_id of
Walther@60729
  2154
        {model = model_patt, program = program as Rule.Prog prog, prog_rls, where_, where_rls,...}=>
Walther@60729
  2155
          (model_patt, program, prog, prog_rls, where_, where_rls)
Walther@60729
  2156
Walther@60729
  2157
\<close> ML \<open>
Walther@60729
  2158
    val return_of_max_variant = 
Walther@60729
  2159
           of_max_variant model_patt i_model;
Walther@60729
  2160
\<close> ML \<open>(*///----------- step into of_max_variant --------------------------------------------\\*)
Walther@60729
  2161
(*///----------------- step into of_max_variant --------------------------------------------\\*)
Walther@60729
  2162
"~~~~~ fun of_max_variant , args:"; val (model_patt, i_model) = (model_patt, i_model);
Walther@60733
  2163
\<close> ML \<open>
Walther@60729
  2164
    val all_variants =
Walther@60729
  2165
        map (fn (_, variants, _, _, _) => variants) i_model
Walther@60729
  2166
        |> flat
Walther@60729
  2167
        |> distinct op =
Walther@60729
  2168
    val variants_separated = map (filter_variants' i_model) all_variants
Walther@60729
  2169
    val sums_corr = map (cnt_corrects) variants_separated
Walther@60729
  2170
    val sum_variant_s = arrange_args1 sums_corr (1, all_variants)
Walther@60729
  2171
    val (_, max_variant) = hd (*..crude decision, up to improvement *)
Walther@60729
  2172
      (sort (fn ((i, _), (j, _)) => int_ord (i, j)) sum_variant_s)
Walther@60729
  2173
    val i_model_max =
Walther@60729
  2174
      filter (fn (_, variants, _ , _ ,_) => member (op =) variants max_variant) i_model
Walther@60729
  2175
    val equal_descr_pairs = map (get_equal_descr i_model) model_patt |> flat
Walther@60729
  2176
;
Walther@60733
  2177
\<close> ML \<open>
Walther@60729
  2178
(*+*)if (equal_descr_pairs |> descr_pairs_to_string @{context}) = "[" ^
Walther@60729
  2179
(*defined: in program+model--vvvv--,----------------------------------------- in problem ---vvv*)
Walther@60733
  2180
  "((#Given, (Traegerlaenge, l_l)), (1, [1], true ,#Given, " ^
Walther@60733
  2181
    "(Cor_TEST Traegerlaenge L , pen2str, Position.T))), " ^
Walther@60729
  2182
  "((#Given, (Streckenlast, q_q)), (2, [1], true ,#Given, " ^
Walther@60733
  2183
    "(Cor_TEST Streckenlast q_0 , pen2str, Position.T))), " ^
Walther@60729
  2184
  "((#Given, (FunktionsVariable, fun_var)), (5, [1], true ,#Given, " ^
Walther@60733
  2185
    "(Cor_TEST FunktionsVariable x , pen2str, Position.T))), " ^
Walther@60729
  2186
  "((#Given, (GleichungsVariablen, vs)), (6, [1], true ,#Given, " ^
Walther@60733
  2187
    "(Cor_TEST GleichungsVariablen [c, c_2, c_3, c_4] , pen2str, Position.T))), " ^
Walther@60729
  2188
  "((#Given, (AbleitungBiegelinie, id_der)), (7, [1], true ,#Given, " ^
Walther@60733
  2189
    "(Cor_TEST AbleitungBiegelinie dy , pen2str, Position.T))), " ^
Walther@60733
  2190
  "((#Find, (Biegelinie, b_b)), (3, [1], true ,#Find, " ^
Walther@60733
  2191
    "(Cor_TEST Biegelinie y , pen2str, Position.T))), " ^
Walther@60729
  2192
  "((#Relate, (Randbedingungen, r_b)), (4, [1], true ,#Relate, " ^
Walther@60733
  2193
    "(Cor_TEST Randbedingungen [y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0] , pen2str, Position.T)))]"
Walther@60729
  2194
then () else error "of_max_variant: equal_descr_pairs CHANGED";
Walther@60729
  2195
Walther@60733
  2196
\<close> ML \<open>
Walther@60729
  2197
    val return_make_env_model = make_env_model equal_descr_pairs;
Walther@60729
  2198
\<close> ML \<open> (*////--------- step into make_env_model --------------------------------------------\\*)
Walther@60729
  2199
(*////---------------- step into make_env_model --------------------------------------------\\*)
Walther@60729
  2200
\<close> ML \<open>
Walther@60729
  2201
"~~~~~ fun make_env_model , args:"; val (equal_descr_pairs) = (equal_descr_pairs);
Walther@60729
  2202
\<close> ML \<open>
Walther@60729
  2203
"~~~~~ fun xxx , args:"; val ((_, (_, id)), (_, _, _, _, (feedb, _))) = nth 4 equal_descr_pairs;
Walther@60729
  2204
\<close> ML \<open>
Walther@60729
  2205
"~~~~~ fun mk_env_model , args:"; val (id, (Model_Def.Cor_TEST ((descr, ts), _))) = (id, feedb);
Walther@60729
  2206
\<close> ML \<open>
Walther@60729
  2207
           handle_lists id (descr, ts);
Walther@60729
  2208
"~~~~~ fun handle_lists , args:"; val (id, (descr, ts)) = (id, (descr, ts));
Walther@60729
  2209
\<close> ML \<open>
Walther@60729
  2210
(*+*)val xxx = ts |> UnparseC.terms @{context}
Walther@60729
  2211
\<close> ML \<open>
Walther@60729
  2212
  (*if*) Model_Pattern.is_list_descr descr (*then*);
Walther@60729
  2213
\<close> ML \<open>
Walther@60729
  2214
    (*if*) length ts > 1 (*then*);
Walther@60729
  2215
\<close> ML \<open>
Walther@60729
  2216
      (*if*) TermC.is_list (hd ts) (*then*);
Walther@60729
  2217
\<close> ML \<open>
Walther@60729
  2218
val return_handle_lists_step =
Walther@60729
  2219
  [(id, TermC.list2isalist (Model_Pattern.typ_of_element descr) (map TermC.the_single ts))]
Walther@60729
  2220
\<close> ML \<open>
Walther@60733
  2221
\<close> ML \<open>
Walther@60733
  2222
\<close> ML \<open>
Walther@60733
  2223
return_handle_lists_step
Walther@60733
  2224
\<close> ML \<open>
Walther@60733
  2225
(*+*)val xxx = return_handle_lists_step |> Subst.to_string @{context}
Walther@60733
  2226
\<close> ML \<open>
Walther@60729
  2227
(*+*)val "[\"\n(vs, [c, c_2, c_3, c_4])\"]"
Walther@60729
  2228
  = return_handle_lists_step |> Subst.to_string @{context}
Walther@60729
  2229
(*\\\\---------------- step into make_env_model --------------------------------------------//*)
Walther@60729
  2230
\<close> ML \<open> (*\\\\--------- step into make_env_model --------------------------------------------//*)
Walther@60729
  2231
    val env_model = return_make_env_model
Walther@60729
  2232
(*+*)val "[\"\n(l_l, L)\", \"\n(q_q, q_0)\", \"\n(fun_var, x)\", \"\n(vs, [c, c_2, c_3, c_4])\", \"\n(id_der, dy)\", \"\n(b_b, y)\", \"\n(r_b, [y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0])\"]"
Walther@60729
  2233
  = env_model |> Subst.to_string @{context}
Walther@60729
  2234
\<close> ML \<open> (*|||---------- contine of_max_variant ------------------------------------------------*)
Walther@60729
  2235
(*|||----------------- contine of_max_variant ------------------------------------------------*)
Walther@60729
  2236
    val equal_givens = filter (fn ((m_field, _), _) => m_field = "#Given") equal_descr_pairs
Walther@60729
  2237
    val subst_eval_list = make_envs_preconds equal_givens
Walther@60729
  2238
    val (env_subst, env_eval) = split_list subst_eval_list
Walther@60729
  2239
\<close> ML \<open>
Walther@60729
  2240
(*bieg*)val "[\"\n(l_l, L)\", \"\n(q_q, q_0)\", \"\n(fun_var, x)\", \"\n(vs, [c, c_2, c_3, c_4])\", \"\n(id_der, dy)\", \"\n(b_b, y)\", \"\n(r_b, [y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0])\"]" = env_model |> Subst.to_string @{context}
Walther@60729
  2241
(*bieg*)val "[]" = env_subst |> Subst.to_string @{context}
Walther@60729
  2242
\<close> ML \<open>
Walther@60729
  2243
val return_of_max_variant_step = (i_model_max, env_model, (env_subst, env_eval))
Walther@60729
  2244
\<close> ML \<open>
Walther@60729
  2245
(*+*)if return_of_max_variant_step = return_of_max_variant then () else error "<>"
Walther@60729
  2246
\<close> ML \<open>(*\\\----------- step into of_max_variant --------------------------------------------//*)
Walther@60729
  2247
(*\\\----------------- step into of_max_variant --------------------------------------------//*)
Walther@60729
  2248
    val (_, env_model, (env_subst, env_eval)) = return_of_max_variant
Walther@60729
  2249
\<close> ML \<open>
Walther@60729
  2250
(*|------------------- contine init_pstate ---------------------------------------------------*)
Walther@60729
  2251
    val actuals = map snd env_model
Walther@60729
  2252
(*+*) val "[L, q_0, x, [c, c_2, c_3, c_4], dy, y, [y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0]]"
Walther@60729
  2253
  = actuals |> UnparseC.terms @{context}
Walther@60729
  2254
\<close> ML \<open>
Walther@60729
  2255
    val formals = Program.formal_args prog
Walther@60729
  2256
(*+*)val "[l_l, q_q, fun_var, b_b, r_b, vs, id_der]"
Walther@60729
  2257
  = formals |> UnparseC.terms @{context}
Walther@60729
  2258
(*+*)val 7 = length formals
Walther@60729
  2259
Walther@60729
  2260
\<close> ML \<open>
Walther@60732
  2261
    val preconds =
Walther@60732
  2262
      Pre_Conds.check_envs_TEST ctxt where_rls where_ (env_model, (env_subst, env_eval))
Walther@60729
  2263
\<close> ML \<open>
Walther@60729
  2264
    val ctxt = ctxt |> ContextC.insert_assumptions (map snd (snd preconds));
Walther@60729
  2265
    val ist = Istate.e_pstate
Walther@60729
  2266
      |> Istate.set_eval prog_rls
Walther@60729
  2267
      |> Istate.set_env_true (relate_args [] formals actuals ctxt prog met_id formals actuals)
Walther@60729
  2268
\<close> ML \<open>
Walther@60729
  2269
(*+*)(relate_args [] formals actuals ctxt prog met_id formals actuals)
Walther@60729
  2270
\<close> ML \<open>
Walther@60729
  2271
val return_init_pstate = (Istate.Pstate ist, ctxt, program)
Walther@60729
  2272
\<close> ML \<open>(*\\------------ step into init_pstate -----------------------------------------------//*)
Walther@60729
  2273
(*\\------------------ step into init_pstate -----------------------------------------------//*)
Walther@60729
  2274
\<close> ML \<open>(*\------------- step into me_Add_Given_AbleitungBieg --------------------------------//*)
Walther@60729
  2275
(*\------------------- step into me_Add_Given_AbleitungBieg --------------------------------//*)
Walther@60729
  2276
\<close> ML \<open> (*\------------ step into me_Add_Given_AbleitungBieg --------------------------------//*)
Walther@60729
  2277
val (p,_,f,nxt,_,pt) = return_Add_Given_AbleitungBieg;
Walther@60729
  2278
                                                 val Apply_Method ["IntegrierenUndKonstanteBestimmen2"] = nxt
Walther@60729
  2279
Walther@60729
  2280
\<close> ML \<open>
Walther@60729
  2281
(* final test ... ----------------------------------------------------------------------------*)
Walther@60729
  2282
(*+*)val ([], Met) = p
Walther@60729
  2283
(*+*)val Apply_Method ["IntegrierenUndKonstanteBestimmen2"] = nxt
Walther@60729
  2284
Walther@60729
  2285
\<close> ML \<open>
Walther@60729
  2286
\<close>
Walther@60729
  2287
Walther@60732
  2288
section \<open>=====eqsy refine_ori' [0 = - 1 * c_4 / - 1, 0 =\n (- 24 * c_4.., ..] ==================\<close>
Walther@60729
  2289
ML \<open> (*\<longrightarrow> refine.sml*)
Walther@60729
  2290
"----------- refine_ori' [0 = - 1 * c_4 / - 1, 0 =\n (- 24 * c_4.., ..] ------------------------";
Walther@60729
  2291
"----------- refine_ori' [0 = - 1 * c_4 / - 1, 0 =\n (- 24 * c_4.., ..] ------------------------";
Walther@60729
  2292
"----------- refine_ori' [0 = - 1 * c_4 / - 1, 0 =\n (- 24 * c_4.., ..] ------------------------";
Walther@60730
  2293
(*overwrites NEW funs in Test_Theory ABOVE*)
Walther@60729
  2294
open Refine
Walther@60729
  2295
open M_Match
Walther@60729
  2296
open Pre_Conds
Walther@60730
  2297
(*overwrites NEW funs in Test_Theory ABOVE*)
Walther@60729
  2298
(*
Walther@60729
  2299
  refinement of the first eqsystem in Biegelinie, IntegrierenUndKonstanteBestimmen2;
Walther@60729
  2300
  this example was the demonstrator;
Walther@60729
  2301
  respective data are in ~/proto4/../exp_Statics_Biegel_Timischl_7-70.xml.
Walther@60729
  2302
  Data for this test are from biegelinie-5.sml, --- me IntegrierenUndKonstanteBestimmen2 ALL ---
Walther@60729
  2303
*)
Walther@60729
  2304
\<close> ML \<open>
Walther@60729
  2305
(*+*)val o_model = [
Walther@60729
  2306
(1, [1], "#Given", @{term "equalities"},
Walther@60729
  2307
  [@{term "[(0::real) = - 1 * c_4 / - 1]"},
Walther@60729
  2308
   @{term "[(0::real) = (- 24 * c_4 + - 24 * L * c_3 + 12 * L \<up> 2 * c_2 + 4 * L \<up> 3 * c +  - 1 * L \<up> 4 * q_0) / - 24]"},
Walther@60729
  2309
   @{term "[(0::real) = c_2]"},
Walther@60729
  2310
   @{term "[(0::real) = (2 * c_2 + 2 * L * c + - 1 * L \<up> 2 * q_0) / 2]"}] ),
Walther@60729
  2311
(*Const ("List.list.Cons", "real \<Rightarrow> real list \<Rightarrow> real list") $ Free ("c", "real") $ Const ("List.list.Nil", "real list"),*)
Walther@60729
  2312
(2, [1], "#Given", @{term "solveForVars"},
Walther@60729
  2313
  [@{term "[c]  ::real list"},
Walther@60729
  2314
   @{term "[c_2]::real list"},
Walther@60729
  2315
   @{term "[c_3]::real list"},
Walther@60729
  2316
   @{term "[c_4]::real list"}] ),
Walther@60729
  2317
(3, [1], "#Find",  @{term "solution"},  [@{term "ss'''::bool list"}] )]
Walther@60730
  2318
;
Walther@60729
  2319
val SOME ["normalise", "4x4", "LINEAR", "system"] =
Walther@60729
  2320
           refine_ori @{context} o_model ["LINEAR", "system"];
Walther@60729
  2321
\<close> ML \<open>
Walther@60729
  2322
"~~~~~ fun refine_ori , args:"; val (ctxt, oris, pblID) =
Walther@60729
  2323
  (@{context}, o_model, ["LINEAR", "system"]);
Walther@60730
  2324
Walther@60729
  2325
    val opt as SOME ["system", "LINEAR", "4x4", "normalise"]
Walther@60730
  2326
    = app_ptyp (refin ctxt ((rev o tl) pblID) oris) pblID (rev pblID);
Walther@60729
  2327
\<close> ML \<open>
Walther@60729
  2328
(*
Walther@60729
  2329
  app_ptyp works strangely, parameter passing is awkward.
Walther@60729
  2330
  Present refin knows structure of Store.T, thus we bypass app_ptyp
Walther@60730
  2331
*)
Walther@60729
  2332
(*!*)val pblID = ["LINEAR", "system"];(**)
Walther@60730
  2333
val return_refin as SOME ["system", "LINEAR", "LINEAR", "4x4", "normalise"] = (**)
Walther@60729
  2334
           refin ctxt (rev pblID) oris (Store.get_node (rev pblID) pblID (get_pbls ()));
Walther@60729
  2335
\<close> ML \<open>
Walther@60729
  2336
"~~~~~ fun refin , args:"; val (ctxt, pblRD, ori, (Store.Node (pI, [py], pys)))
Walther@60729
  2337
  = (ctxt, (rev pblID), oris, Store.get_node (rev pblID) pblID (get_pbls ()));
Walther@60729
  2338
\<close> ML \<open>
Walther@60729
  2339
      val {where_rls, model, where_, ...} = py: Problem.T
Walther@60729
  2340
      val model = map (Model_Pattern.adapt_to_type ctxt) model
Walther@60730
  2341
      val where_ = map (ParseC.adapt_term_to_type ctxt) where_;
Walther@60730
  2342
\<close> ML \<open>
Walther@60730
  2343
      (*if*) M_Match.match_oris ctxt where_rls ori (model, where_) (*then*);
Walther@60730
  2344
\<close> ML \<open>
Walther@60730
  2345
val return_refin as SOME ["system", "LINEAR"] = SOME (pblRD (** )@ [pI]( **))
Walther@60730
  2346
\<close> ML \<open>
Walther@60730
  2347
(*follow the trace created in parallel by writeln in Store.get_node and Refine.refin*)
Walther@60730
  2348
(*!*)val pblID = ["4x4", "LINEAR", "system"];(**)
Walther@60730
  2349
val return_refin as SOME ["system", "LINEAR", "4x4", "4x4", "normalise"] =
Walther@60730
  2350
          refin ctxt (rev pblID) oris (Store.get_node (rev pblID) pblID (get_pbls ()));
Walther@60730
  2351
\<close> ML \<open>
Walther@60730
  2352
"~~~~~ fun refin , args:"; val (ctxt, pblRD, ori, (Store.Node (pI, [py], pys)))
Walther@60730
  2353
  = (ctxt, (rev pblID), oris, Store.get_node (rev pblID) pblID (get_pbls ()));
Walther@60730
  2354
\<close> ML \<open>
Walther@60730
  2355
      val {where_rls, model, where_, ...} = py: Problem.T
Walther@60730
  2356
      val model = map (Model_Pattern.adapt_to_type ctxt) model
Walther@60729
  2357
      val where_ = map (ParseC.adapt_term_to_type ctxt) where_
Walther@60729
  2358
\<close> ML \<open>
Walther@60730
  2359
(*+*)val (**)true(** )false( **) = (*if*)
Walther@60730
  2360
   M_Match.match_oris ctxt where_rls o_model (model, where_);
Walther@60730
  2361
"~~~~~ fun match_oris , args:"; val (ctxt, where_rls, o_model, (model_pattern, where_)) =
Walther@60730
  2362
  (ctxt, where_rls, o_model, (model, where_));
Walther@60730
  2363
    val i_model = (flat o (map (chk1_ model_pattern))) o_model;
Walther@60730
  2364
Walther@60729
  2365
(*+*)val "[\n(1, [\"1\"], #Given, equalities, [\"[0 = - 1 * c_4 / - 1]\", \"[0 =\n (- 24 * c_4 + - 24 * L * c_3 + 12 * L \<up> 2 * c_2 + 4 * L \<up> 3 * c +\n  - 1 * L \<up> 4 * q_0) /\n - 24]\", \"[0 = c_2]\", \"[0 = (2 * c_2 + 2 * L * c + - 1 * L \<up> 2 * q_0) / 2]\"]), \n(2, [\"1\"], #Given, solveForVars, [\"[c]\", \"[c_2]\", \"[c_3]\", \"[c_4]\"]), \n(3, [\"1\"], #Find, solution, [\"ss'''\"])]"
Walther@60730
  2366
  = o_model |> O_Model.to_string @{context};
Walther@60729
  2367
(*+*)if "[\n" ^
Walther@60729
  2368
  "(1 ,[1] ,true ,#Given ,Cor equalities\n [0 = - 1 * c_4 / - 1,\n  0 =\n  (- 24 * c_4 + - 24 * L * c_3 + 12 * L \<up> 2 * c_2 + 4 * L \<up> 3 * c +\n   - 1 * L \<up> 4 * q_0) /\n  - 24,\n  0 = c_2, 0 = (2 * c_2 + 2 * L * c + - 1 * L \<up> 2 * q_0) / 2] ,(e_s, [[0 = - 1 * c_4 / - 1,\n 0 =\n (- 24 * c_4 + - 24 * L * c_3 + 12 * L \<up> 2 * c_2 + 4 * L \<up> 3 * c +\n  - 1 * L \<up> 4 * q_0) /\n - 24,\n 0 = c_2, 0 = (2 * c_2 + 2 * L * c + - 1 * L \<up> 2 * q_0) / 2]])), \n" ^
Walther@60729
  2369
  "(2 ,[1] ,true ,#Given ,Cor solveForVars [c, c_2, c_3, c_4] ,(v_s, [[c, c_2, c_3, c_4]])), \n" ^
Walther@60729
  2370
  "(3 ,[1] ,true ,#Find ,Cor solution ss''' ,(ss''', [ss''']))]"
Walther@60729
  2371
 = (i_model |> I_Model.to_string @{context}) then () else raise error "i_model CAHNGED";
Walther@60729
  2372
\<close> ML \<open>
Walther@60730
  2373
(*+*)val "[\"(#Given, (equalities, e_s))\", \"(#Given, (solveForVars, v_s))\", \"(#Find, (solution, ss'''))\"]"
Walther@60730
  2374
  = model_pattern |> Model_Pattern.to_string @{context}
Walther@60730
  2375
(*+*)val "[Length e_s = 4, Length v_s = 4]" = where_ |> UnparseC.terms @{context}
Walther@60730
  2376
Walther@60733
  2377
    val mvat = I_Model.variables_TEST model_pattern (I_Model.OLD_to_TEST i_model)
Walther@60730
  2378
    val complete = chk1_mis mvat i_model model_pattern;
Walther@60730
  2379
Walther@60730
  2380
    val (pb as true, (** )_( **)where_'(**)) =
Walther@60730
  2381
 Pre_Conds.check_OLD ctxt where_rls where_ (model_pattern, I_Model.OLD_to_TEST i_model);
Walther@60730
  2382
"~~~~~ fun check_OLD , args:"; val (ctxt, where_rls, pres, (model_patt, i_model)) =
Walther@60730
  2383
  (ctxt, where_rls, where_, (model_pattern, I_Model.OLD_to_TEST i_model));
Walther@60729
  2384
\<close> ML \<open>
Walther@60729
  2385
(*+*)val [(true, xxx), (true, yyy)] = where_'
Walther@60729
  2386
\<close> ML \<open>
Walther@60729
  2387
(*+*)val "[Length\n [0 = - 1 * c_4 / - 1,\n  0 =\n  (- 24 * c_4 + - 24 * L * c_3 + 12 * L \<up> 2 * c_2 + 4 * L \<up> 3 * c +\n   - 1 * L \<up> 4 * q_0) /\n  - 24,\n  0 = c_2, 0 = (2 * c_2 + 2 * L * c + - 1 * L \<up> 2 * q_0) / 2] =\n4, Length [c, c_2, c_3, c_4] = 4]"
Walther@60729
  2388
 = where_' |> map #2 |> UnparseC.terms @{context}
Walther@60730
  2389
Walther@60729
  2390
      val (_, env_model, (env_subst, env_eval)) =
Walther@60730
  2391
           of_max_variant model_patt i_model;
Walther@60729
  2392
"~~~~~ fun of_max_variant , args:"; val (model_patt, i_model) = (model_patt, i_model);
Walther@60729
  2393
\<close> ML \<open>
Walther@60729
  2394
    val all_variants =
Walther@60729
  2395
        map (fn (_, variants, _, _, _) => variants) i_model
Walther@60729
  2396
        |> flat
Walther@60729
  2397
        |> distinct op =
Walther@60729
  2398
    val variants_separated = map (filter_variants' i_model) all_variants
Walther@60729
  2399
    val sums_corr = map (cnt_corrects) variants_separated
Walther@60729
  2400
    val sum_variant_s = arrange_args1 sums_corr (1, all_variants)
Walther@60729
  2401
    val (_, max_variant) = hd (*..crude decision, up to improvement *)
Walther@60729
  2402
      (sort (fn ((i, _), (j, _)) => int_ord (i, j)) sum_variant_s)
Walther@60729
  2403
    val i_model_max =
Walther@60729
  2404
      filter (fn (_, variants, _ , _ ,_) => member (op =) variants max_variant) i_model
Walther@60729
  2405
    val equal_descr_pairs = map (get_equal_descr i_model) model_patt |> flat
Walther@60729
  2406
Walther@60729
  2407
    val env_model = make_env_model equal_descr_pairs
Walther@60729
  2408
    val equal_givens = filter (fn ((m_field, _), _) => m_field = "#Given") equal_descr_pairs
Walther@60729
  2409
\<close> ML \<open>
Walther@60729
  2410
    val subst_eval_list =
Walther@60730
  2411
           make_envs_preconds equal_givens;
Walther@60729
  2412
\<close> ML \<open>
Walther@60729
  2413
"~~~~~ fun make_envs_preconds , args:"; val (equal_givens) = (equal_givens);
Walther@60729
  2414
\<close> ML \<open>
Walther@60729
  2415
"~~~~~ fun xxx , args:"; val ((_, (_, id)), (_, _, _, _, (feedb, _))) = (nth 1 equal_givens);
Walther@60729
  2416
\<close> ML \<open>
Walther@60729
  2417
"~~~~~ fun discern_feedback , args:"; val (id, (Model_Def.Cor_TEST ((descr, ts), _))) =
Walther@60729
  2418
  (id, feedb);
Walther@60729
  2419
\<close> ML \<open>
Walther@60729
  2420
val return_discern_typ as [] =
Walther@60729
  2421
           discern_typ id (descr, ts);
Walther@60729
  2422
\<close> ML \<open>
Walther@60729
  2423
"~~~~~ fun discern_typ , args:"; val (id, (descr, ts)) = (id, (descr, ts));
Walther@60729
  2424
\<close> ML \<open>
Walther@60729
  2425
(*+*)val "[[0 = - 1 * c_4 / - 1], [0 =\n (- 24 * c_4 + - 24 * L * c_3 + 12 * L \<up> 2 * c_2 + 4 * L \<up> 3 * c +\n  - 1 * L \<up> 4 * q_0) /\n - 24], [0 = c_2], [0 = (2 * c_2 + 2 * L * c + - 1 * L \<up> 2 * q_0) / 2]]"
Walther@60729
  2426
 = ts |> UnparseC.terms @{context}
Walther@60729
  2427
\<close> ML \<open>
Walther@60730
  2428
      val ts = if Model_Pattern.is_list_descr descr
Walther@60730
  2429
        then if TermC.is_list (hd ts)
Walther@60730
  2430
          then ts |> map TermC.isalist2list |> flat
Walther@60730
  2431
          else ts
Walther@60730
  2432
        else ts
Walther@60730
  2433
Walther@60730
  2434
(*+*)val "[0 = - 1 * c_4 / - 1, 0 =\n(- 24 * c_4 + - 24 * L * c_3 + 12 * L \<up> 2 * c_2 + 4 * L \<up> 3 * c +\n - 1 * L \<up> 4 * q_0) /\n- 24, 0 = c_2, 0 = (2 * c_2 + 2 * L * c + - 1 * L \<up> 2 * q_0) / 2]"
Walther@60730
  2435
 = ts |> UnparseC.terms @{context};
Walther@60730
  2436
\<close> ML \<open>
Walther@60729
  2437
  (*if*)  Model_Pattern.typ_of_element descr = HOLogic.boolT (*then*);
Walther@60729
  2438
\<close> ML \<open>
Walther@60729
  2439
(*+*)val false = all_lhs_atoms ts
Walther@60729
  2440
(*-------------------- contine check_OLD -----------------------------------------------------*)
Walther@60729
  2441
 
Walther@60730
  2442
     val pres_subst = map (TermC.subst_atomic_all env_subst) where_;
Walther@60729
  2443
\<close> ML \<open>
Walther@60729
  2444
(*+*)if "[\"\n" ^
Walther@60729
  2445
  "(e_s, [0 = - 1 * c_4 / - 1,\n 0 =\n (- 24 * c_4 + - 24 * L * c_3 + 12 * L \<up> 2 * c_2 + 4 * L \<up> 3 * c +\n  - 1 * L \<up> 4 * q_0) /\n - 24,\n 0 = c_2, 0 = (2 * c_2 + 2 * L * c + - 1 * L \<up> 2 * q_0) / 2])\", \"\n" ^
Walther@60729
  2446
  "(v_s, [c, c_2, c_3, c_4])\", \"\n(ss''', ss''')\"]"
Walther@60729
  2447
 = (env_model |> Subst.to_string @{context}) then () else error "env_model CAHNGED";
Walther@60729
  2448
\<close> ML \<open>
Walther@60729
  2449
(*+*)val "[]" = env_eval |> Subst.to_string @{context} (*GOON \<rightarrow> []*)
Walther@60729
  2450
\<close> ML \<open>
Walther@60729
  2451
      val pres_subst_other = map (TermC.subst_atomic_all env_model) (map #2 pres_subst);
Walther@60729
  2452
\<close> ML \<open>
Walther@60729
  2453
      val full_subst = if env_eval = [] then pres_subst_other
Walther@60729
  2454
        else map (TermC.subst_atomic_all env_eval) (map #2 pres_subst_other)
Walther@60729
  2455
\<close> ML \<open>
Walther@60729
  2456
      val evals = map (eval ctxt where_rls) full_subst
Walther@60729
  2457
\<close> ML \<open>
Walther@60729
  2458
val return_check_OLD = (foldl and_ (true, map fst evals), pres_subst_other)
Walther@60729
  2459
\<close> ML \<open>
Walther@60729
  2460
val return_refin as SOME ["system", "LINEAR", "4x4"] = SOME (pblRD(** ) @ [pI]( **))
Walther@60729
  2461
\<close> ML \<open>
Walther@60729
  2462
(*follow the trace created in parallel by writeln in Store.get_node and Refine.refin*)
Walther@60729
  2463
(*!*)val pblID = ["normalise", "4x4", "LINEAR", "system"];(**)
Walther@60730
  2464
val return_refin as SOME ["system", "LINEAR", "4x4", "normalise", "normalise"] =
Walther@60729
  2465
          refin ctxt (rev pblID) oris (Store.get_node (rev pblID) pblID (get_pbls ()));
Walther@60729
  2466
\<close> ML \<open>
Walther@60729
  2467
\<close> ML \<open>
Walther@60729
  2468
\<close> ML \<open>
Walther@60729
  2469
\<close>
Walther@60729
  2470
Walther@60729
  2471
section \<open>===================================================================================\<close>
Walther@60732
  2472
section \<open>=====isa2 -> me IntegrierenUndKonstanteBestimmen2 ALL ====================================\<close>
Walther@60729
  2473
ML \<open>
Walther@60729
  2474
open Sub_Problem
Walther@60729
  2475
\<close> ML \<open> (*\<rightarrow> biegelinie-4.sml OR NEW biegelinie-5.sml *)
Walther@60729
  2476
"----------- me IntegrierenUndKonstanteBestimmen2 ALL ----------------------------------------";
Walther@60729
  2477
"----------- me IntegrierenUndKonstanteBestimmen2 ALL ----------------------------------------";
Walther@60729
  2478
"----------- me IntegrierenUndKonstanteBestimmen2 ALL ----------------------------------------";
Walther@60729
  2479
(*
Walther@60729
  2480
  This test is deeply nested (down into details of creating environments).
Walther@60729
  2481
  In order to make Sidekick show the structure add to each
Walther@60729
  2482
  *                    (*/...\*) and        (*\.../*)
Walther@60729
  2483
  * a companion  > ML < (*/...\*) and > ML < (*\.../*)
Walther@60729
  2484
  Note the wrong ^----^ delimiters.
Walther@60729
  2485
*)
Walther@60729
  2486
val fmz = ["Traegerlaenge L", "Streckenlast q_0", "Biegelinie y",
Walther@60729
  2487
	     "Randbedingungen [y 0 = (0::real), y L = 0, M_b 0 = 0, M_b L = 0]",
Walther@60729
  2488
	     "FunktionsVariable x", "GleichungsVariablen [c, c_2, c_3, c_4]",
Walther@60729
  2489
	     "AbleitungBiegelinie dy"];
Walther@60729
  2490
val (dI',pI',mI') = ("Biegelinie", ["Biegelinien"],
Walther@60729
  2491
		     ["IntegrierenUndKonstanteBestimmen2"]);
Walther@60729
  2492
val p = e_pos'; val c = [];
Walther@60729
  2493
(*[], Pbl*)val (p,_,f,nxt,_,pt) = Test_Code.init_calc @{context} [(fmz, (dI',pI',mI'))]; val Model_Problem = nxt
Walther@60729
  2494
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "Traegerlaenge L" = nxt
Walther@60729
  2495
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "Streckenlast q_0" = nxt
Walther@60729
  2496
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "Biegelinie y" = nxt
Walther@60729
  2497
(*/---broken elementwise input to lists---\* )
Walther@60729
  2498
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; (*\<rightarrow>Add_Relation "Randbedingungen [y 0 = 0]", ERROR MISSING step: M_b 0 = 0*)
Walther@60729
  2499
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; (*\<rightarrow>Add_Relation "Randbedingungen [y L = 0, M_b 0 = 0, M_b L = 0]"*)
Walther@60729
  2500
( *\---broken elementwise input to lists---/*)
Walther@60729
  2501
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Relation "Randbedingungen [y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0]" = nxt
Walther@60729
  2502
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2503
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["Biegelinien"] = nxt
Walther@60729
  2504
(*[], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["IntegrierenUndKonstanteBestimmen2"] = nxt
Walther@60729
  2505
(*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "FunktionsVariable x" = nxt
Walther@60729
  2506
(*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "GleichungsVariablen [c, c_2, c_3, c_4]" = nxt
Walther@60729
  2507
(*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "AbleitungBiegelinie dy" = nxt
Walther@60729
  2508
(*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["IntegrierenUndKonstanteBestimmen2"] = nxt
Walther@60729
  2509
\<close> ML \<open>
Walther@60729
  2510
(*[1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2511
(*[1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "Streckenlast q_0" = nxt
Walther@60729
  2512
(*[1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "FunktionsVariable x" = nxt
Walther@60729
  2513
(*[1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "Funktionen funs'''" = nxt
Walther@60729
  2514
(*[1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2515
(*[1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["vonBelastungZu", "Biegelinien"] = nxt
Walther@60729
  2516
(*[1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["Biegelinien", "ausBelastung"] = nxt
Walther@60729
  2517
(*[1], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "Biegelinie y" = nxt
Walther@60729
  2518
(*[1], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "AbleitungBiegelinie dy" = nxt
Walther@60729
  2519
(*[1], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["Biegelinien", "ausBelastung"] = nxt
Walther@60729
  2520
\<close> ML \<open>
Walther@60729
  2521
(*[1, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite ("sym_neg_equal_iff_equal", _) = nxt
Walther@60729
  2522
(*[1, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite ("Belastung_Querkraft", _) = nxt
Walther@60729
  2523
(*[1, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Subproblem ("Biegelinie", ["named", "integrate", "function"]) = nxt
Walther@60729
  2524
(*[1, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2525
(*[1, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "functionTerm (- q_0)" = nxt
Walther@60729
  2526
(*[1, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "integrateBy x" = nxt
Walther@60729
  2527
(*[1, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "antiDerivativeName Q" = nxt
Walther@60729
  2528
(*[1, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2529
(*[1, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["named", "integrate", "function"] = nxt
Walther@60729
  2530
(*[1, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["diff", "integration", "named"] = nxt
Walther@60729
  2531
(*[1, 3], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["diff", "integration", "named"] = nxt
Walther@60729
  2532
\<close> ML \<open>
Walther@60729
  2533
(*[1, 3, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set_Inst (["(''bdv'', x)"], "simplify_Integral") = nxt
Walther@60729
  2534
(*[1, 3, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set_Inst (["(''bdv'', x)"], "integration") = nxt
Walther@60729
  2535
(*[1, 3, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["named", "integrate", "function"] = nxt
Walther@60729
  2536
\<close> ML \<open>
Walther@60729
  2537
(*[1, 3], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite ("Querkraft_Moment", _) = nxt
Walther@60729
  2538
(*[1, 4], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Subproblem ("Biegelinie", ["named", "integrate", "function"]) = nxt
Walther@60729
  2539
(*[1, 5], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2540
(*[1, 5], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "functionTerm (c + - 1 * q_0 * x)" = nxt
Walther@60729
  2541
(*[1, 5], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "integrateBy x" = nxt
Walther@60729
  2542
(*[1, 5], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "antiDerivativeName M_b" = nxt
Walther@60729
  2543
(*[1, 5], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2544
(*[1, 5], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["named", "integrate", "function"] = nxt
Walther@60729
  2545
(*[1, 5], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["diff", "integration", "named"] = nxt
Walther@60729
  2546
(*[1, 5], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["diff", "integration", "named"] = nxt
Walther@60729
  2547
\<close> ML \<open>
Walther@60729
  2548
(*[1, 5, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set_Inst (["(''bdv'', x)"], "integration") = nxt
Walther@60729
  2549
(*[1, 5, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["named", "integrate", "function"] = nxt
Walther@60729
  2550
\<close> ML \<open>
Walther@60729
  2551
(*[1, 5], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite ("Moment_Neigung", _) = nxt
Walther@60729
  2552
(*[1, 6], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite ("make_fun_explicit", _) = nxt
Walther@60729
  2553
(*[1, 7], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Subproblem ("Biegelinie", ["named", "integrate", "function"]) = nxt
Walther@60729
  2554
(*[1, 8], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2555
(*[1, 8], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "functionTerm (1 / - EI * (c_2 + c * x + - 1 * q_0 / 2 * x \<up> 2))" = nxt
Walther@60729
  2556
(*[1, 8], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "integrateBy x" = nxt
Walther@60729
  2557
(*[1, 8], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "antiDerivativeName dy" = nxt
Walther@60729
  2558
(*[1, 8], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2559
(*[1, 8], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["named", "integrate", "function"] = nxt
Walther@60729
  2560
(*[1, 8], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["diff", "integration", "named"] = nxt
Walther@60729
  2561
(*[1, 8], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["diff", "integration", "named"] = nxt
Walther@60729
  2562
\<close> ML \<open>
Walther@60729
  2563
(*[1, 8, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set_Inst (["(''bdv'', x)"], "simplify_Integral") = nxt
Walther@60729
  2564
(*[1, 8, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set_Inst (["(''bdv'', x)"], "integration") = nxt
Walther@60729
  2565
(*[1, 8, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["named", "integrate", "function"] = nxt
Walther@60729
  2566
\<close> ML \<open>
Walther@60729
  2567
(*[1, 8], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Subproblem ("Biegelinie", ["named", "integrate", "function"]) = nxt
Walther@60729
  2568
(*[1, 9], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2569
(*[1, 9], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "functionTerm\n (c_3 +\n  1 / (- 1 * EI) *\n  (c_2 * x + c / 2 * x \<up> 2 + - 1 * q_0 / 6 * x \<up> 3))" = nxt
Walther@60729
  2570
(*[1, 9], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "integrateBy x" = nxt
Walther@60729
  2571
(*[1, 9], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "antiDerivativeName y" = nxt
Walther@60729
  2572
(*[1, 9], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2573
(*[1, 9], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["named", "integrate", "function"] = nxt
Walther@60729
  2574
(*[1, 9], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["diff", "integration", "named"] = nxt
Walther@60729
  2575
(*[1, 9], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["diff", "integration", "named"] = nxt
Walther@60729
  2576
\<close> ML \<open>
Walther@60729
  2577
(*[1, 9, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set_Inst (["(''bdv'', x)"], "integration") = nxt
Walther@60729
  2578
(*[1, 9, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["named", "integrate", "function"] = nxt
Walther@60729
  2579
(*[1, 9], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["vonBelastungZu", "Biegelinien"] = nxt
Walther@60729
  2580
(*[1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Subproblem ("Biegelinie", ["setzeRandbedingungen", "Biegelinien"]) = nxt
Walther@60729
  2581
\<close> ML \<open>
Walther@60729
  2582
(*[2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2583
(*[2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "Funktionen\n [Q x = c + - 1 * q_0 * x, M_b x = c_2 + c * x + - 1 * q_0 / 2 * x \<up> 2,\n  dy x =\n  c_3 +\n  1 / (- 1 * EI) *\n  (c_2 * x + c / 2 * x \<up> 2 + - 1 * q_0 / 6 * x \<up> 3),\n  y x =\n  c_4 + c_3 * x +\n  1 / (- 1 * EI) *\n  (c_2 / 2 * x \<up> 2 + c / 6 * x \<up> 3 + - 1 * q_0 / 24 * x \<up> 4)]" = nxt
Walther@60729
  2584
(*/---broken elementwise input to lists---\* )
Walther@60729
  2585
(*[2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; (*\<rightarrow>Add_Relation "Randbedingungen [y 0 = 0]", ERROR MISSING step: M_b 0 = 0*)
Walther@60729
  2586
(*[2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; (*\<rightarrow>Add_Relation "Randbedingungen [y L = 0, M_b 0 = 0, M_b L = 0]"*)
Walther@60729
  2587
( *\---broken elementwise input to lists---/*)
Walther@60729
  2588
(*[2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "Randbedingungen [y 0 = 0, y L = 0, M_b 0 = 0, M_b L = 0]" = nxt
Walther@60729
  2589
(*[2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "Gleichungen equs'''" = nxt
Walther@60729
  2590
(*[2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2591
(*[2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["setzeRandbedingungen", "Biegelinien"] = nxt
Walther@60729
  2592
(*[2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["Biegelinien", "setzeRandbedingungenEin"] = nxt
Walther@60729
  2593
(*[2], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["Biegelinien", "setzeRandbedingungenEin"] = nxt
Walther@60729
  2594
\<close> ML \<open>
Walther@60729
  2595
(*[2, 1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2596
(*[2, 1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "functionEq\n (y x =\n  c_4 + c_3 * x +\n  1 / (- 1 * EI) *\n  (c_2 / 2 * x \<up> 2 + c / 6 * x \<up> 3 + - 1 * q_0 / 24 * x \<up> 4))" = nxt
Walther@60729
  2597
(*[2, 1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val  Add_Given "substitution (y 0 = 0)" = nxt
Walther@60729
  2598
(*[2, 1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "equality equ'''" = nxt
Walther@60729
  2599
(*[2, 1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2600
(*[2, 1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["makeFunctionTo", "equation"] = nxt
Walther@60729
  2601
(*[2, 1], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["Equation", "fromFunction"] = nxt
Walther@60729
  2602
(*[2, 1], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["Equation", "fromFunction"] = nxt
Walther@60729
  2603
\<close> ML \<open>
Walther@60729
  2604
(*[2, 1, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Substitute ["x = 0"] = nxt
Walther@60729
  2605
(*[2, 1, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Substitute ["y 0 = 0"] = nxt
Walther@60729
  2606
(*[2, 1, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set "norm_Rational" = nxt
Walther@60729
  2607
(*[2, 1, 3], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["makeFunctionTo", "equation"] = nxt
Walther@60729
  2608
\<close> ML \<open>
Walther@60729
  2609
(*[2, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Subproblem ("Biegelinie", ["makeFunctionTo", "equation"]) = nxt
Walther@60729
  2610
(*[2, 2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2611
(*[2, 2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "functionEq\n (y x =\n  c_4 + c_3 * x +\n  1 / (- 1 * EI) *\n  (c_2 / 2 * x \<up> 2 + c / 6 * x \<up> 3 + - 1 * q_0 / 24 * x \<up> 4))" = nxt
Walther@60729
  2612
(*[2, 2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "substitution (y L = 0)" = nxt
Walther@60729
  2613
(*[2, 2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "equality equ'''" = nxt
Walther@60729
  2614
(*[2, 2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2615
(*[2, 2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["makeFunctionTo", "equation"] = nxt
Walther@60729
  2616
(*[2, 2], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["Equation", "fromFunction"] = nxt
Walther@60729
  2617
(*[2, 2], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["Equation", "fromFunction"] = nxt
Walther@60729
  2618
\<close> ML \<open>
Walther@60729
  2619
(*[2, 2, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Substitute ["x = L"] = nxt
Walther@60729
  2620
(*[2, 2, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Substitute ["y L = 0"] = nxt
Walther@60729
  2621
(*[2, 2, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set "norm_Rational" = nxt
Walther@60729
  2622
(*[2, 2, 3], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["makeFunctionTo", "equation"] = nxt
Walther@60729
  2623
\<close> ML \<open>
Walther@60729
  2624
(*[2, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Subproblem ("Biegelinie", ["makeFunctionTo", "equation"]) = nxt
Walther@60729
  2625
(*[2, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2626
(*[2, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "functionEq (M_b x = c_2 + c * x + - 1 * q_0 / 2 * x \<up> 2)" = nxt
Walther@60729
  2627
(*[2, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "substitution (M_b 0 = 0)" = nxt
Walther@60729
  2628
(*[2, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "equality equ'''" = nxt
Walther@60729
  2629
(*[2, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2630
(*[2, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["makeFunctionTo", "equation"] = nxt
Walther@60729
  2631
(*[2, 3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["Equation", "fromFunction"] = nxt
Walther@60729
  2632
(*[2, 3], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["Equation", "fromFunction"] = nxt
Walther@60729
  2633
\<close> ML \<open>
Walther@60729
  2634
(*[2, 3, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Substitute ["x = 0"] = nxt
Walther@60729
  2635
(*[2, 3, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Substitute ["M_b 0 = 0"] = nxt
Walther@60729
  2636
(*[2, 3, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set "norm_Rational" = nxt
Walther@60729
  2637
(*[2, 3, 3], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["makeFunctionTo", "equation"] = nxt
Walther@60729
  2638
\<close> ML \<open>
Walther@60729
  2639
(*[2, 3], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Subproblem ("Biegelinie", ["makeFunctionTo", "equation"]) = nxt
Walther@60729
  2640
(*[2, 4], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Model_Problem = nxt
Walther@60729
  2641
(*[2, 4], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "functionEq (M_b x = c_2 + c * x + - 1 * q_0 / 2 * x \<up> 2)" = nxt
Walther@60729
  2642
(*[2, 4], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Given "substitution (M_b L = 0)" = nxt
Walther@60729
  2643
(*[2, 4], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Add_Find "equality equ'''" = nxt
Walther@60729
  2644
(*[2, 4], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Theory "Biegelinie" = nxt
Walther@60729
  2645
(*[2, 4], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Problem ["makeFunctionTo", "equation"] = nxt
Walther@60729
  2646
(*[2, 4], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Specify_Method ["Equation", "fromFunction"] = nxt
Walther@60729
  2647
(*[2, 4], Met*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Apply_Method ["Equation", "fromFunction"] = nxt
Walther@60729
  2648
\<close> ML \<open>
Walther@60729
  2649
(*[2, 4, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Substitute ["x = L"] = nxt
Walther@60729
  2650
(*[2, 4, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Substitute ["M_b L = 0"] = nxt
Walther@60729
  2651
(*[2, 4, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Rewrite_Set "norm_Rational" = nxt
Walther@60729
  2652
(*[2, 4, 3], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["makeFunctionTo", "equation"] = nxt
Walther@60729
  2653
(*[2, 4], Res*)val (p,_,f,nxt,_,pt) = me nxt p c pt; val Check_Postcond ["setzeRandbedingungen", "Biegelinien"] = nxt
Walther@60729
  2654
Walther@60729
  2655
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2656
(*[2], Res*)val return_Check_Postcond_setzeRandbedingungen = me nxt p c pt;
Walther@60729
  2657
\<close> ML \<open> (*/------------ step into me_Check_Postcond_setzeRandbedingungen --------------------\\*)
Walther@60729
  2658
(*/------------------- step into me_Check_Postcond_setzeRandbedingungen --------------------\\*)
Walther@60729
  2659
"~~~~~ fun me , args:"; val (tac, p, _(*NEW remove*), pt) = (nxt, p, c, pt);
Walther@60729
  2660
\<close> ML \<open>
Walther@60729
  2661
      val ctxt = Ctree.get_ctxt pt p
Walther@60729
  2662
\<close> ML \<open>
Walther@60729
  2663
      val (pt, p) = 
Walther@60729
  2664
  	    case Step.by_tactic tac (pt, p) of
Walther@60729
  2665
  		    ("ok", (_, _, ptp)) => ptp
Walther@60729
  2666
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2667
    (*case*)
Walther@60729
  2668
      Step.do_next p ((pt, Pos.e_pos'), []) (*of*);
Walther@60729
  2669
\<close> ML \<open>
Walther@60729
  2670
"~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis)) = (p, ((pt, Pos.e_pos'), []));
Walther@60729
  2671
\<close> ML \<open>
Walther@60729
  2672
  (*if*) Pos.on_calc_end ip (*else*);
Walther@60729
  2673
      val (_, probl_id, _) = Calc.references (pt, p);
Walther@60729
  2674
\<close> ML \<open>                                                               
Walther@60729
  2675
val _ =                                                              
Walther@60729
  2676
      (*case*) tacis (*of*);
Walther@60729
  2677
\<close> ML \<open>
Walther@60729
  2678
        (*if*) probl_id = Problem.id_empty (*else*);
Walther@60729
  2679
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2680
           switch_specify_solve p_ (pt, ip);
Walther@60729
  2681
\<close> ML \<open>
Walther@60729
  2682
"~~~~~ fun switch_specify_solve , args:"; val (state_pos, (pt, input_pos)) = (p_, (pt, ip));
Walther@60729
  2683
\<close> ML \<open>
Walther@60729
  2684
      (*if*) Pos.on_specification ([], state_pos) (*else*);
Walther@60729
  2685
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2686
        LI.do_next (pt, input_pos)
Walther@60729
  2687
\<close> ML \<open>
Walther@60729
  2688
"~~~~~ and do_next , args:"; val (ptp as (pt, pos as (p, p_))) = (pt, input_pos);
Walther@60729
  2689
\<close> ML \<open>
Walther@60729
  2690
    (*if*) MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p) (*else*);
Walther@60729
  2691
\<close> ML \<open>
Walther@60729
  2692
	      val ((ist, ctxt), sc) = LItool.resume_prog pos pt;
Walther@60729
  2693
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2694
val Next_Step (Pstate _, _, Subproblem' _) =
Walther@60729
  2695
  (*case*) find_next_step sc ptp ist ctxt (*of*);
Walther@60729
  2696
\<close> ML \<open>
Walther@60729
  2697
"~~~~~ fun find_next_step , args:"; val ((Rule.Prog prog), (ptp as (pt, pos as (p, _))), (Pstate ist), ctxt)
Walther@60729
  2698
  = (sc, ptp, ist, ctxt);
Walther@60729
  2699
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2700
val Accept_Tac (Subproblem' _, _, _) =
Walther@60729
  2701
  (*case*) scan_to_tactic (prog, (ptp, ctxt)) (Pstate ist) (*of*);
Walther@60729
  2702
\<close> ML \<open>
Walther@60729
  2703
"~~~~~ fun scan_to_tactic , args:"; val ((prog, cc), (Pstate (ist as {path, ...})))
Walther@60729
  2704
  = ((prog, (ptp, ctxt)), (Pstate ist));
Walther@60729
  2705
\<close> ML \<open>
Walther@60729
  2706
  (*if*) path = [] (*else*);
Walther@60729
  2707
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2708
val Accept_Tac (Subproblem' _, _, _) =
Walther@60729
  2709
           go_scan_up (prog, cc) (trans_scan_up ist);
Walther@60729
  2710
\<close> ML \<open>
Walther@60729
  2711
"~~~~~ fun go_scan_up , args:"; val ((pcc as (sc, (_, ctxt))), (ist as {path, act_arg, found_accept, ...}))
Walther@60729
  2712
  = ((prog, cc), (trans_scan_up ist));
Walther@60729
  2713
\<close> ML \<open>
Walther@60729
  2714
    (*if*) path = [R] (*root of the program body*) (*else*);
Walther@60729
  2715
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2716
           scan_up pcc (ist |> path_up) (go_up ctxt path sc)
Walther@60729
  2717
\<close> ML \<open>
Walther@60729
  2718
"~~~~~ and scan_up , args:"; val ((pcc as (sc, cc as (_, ctxt))), ist, (Const (\<^const_name>\<open>Let\<close>(*1*), _) $ _))
Walther@60729
  2719
  = (pcc, (ist |> path_up), (go_up ctxt path sc));
Walther@60729
  2720
\<close> ML \<open>
Walther@60729
  2721
        val (i, body) = check_Let_up ctxt ist sc
Walther@60729
  2722
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2723
val Accept_Tac (Subproblem' _, _, _) =
Walther@60729
  2724
  (*case*) scan_dn cc (ist |> path_up_down [R, D] |> upd_env i) body (*of*);
Walther@60729
  2725
\<close> ML \<open>
Walther@60729
  2726
"~~~~~ fun scan_dn , args:"; val (cc, ist, (Const (\<^const_name>\<open>Let\<close>(*1*), _) $ e $ (Abs (i, T, b))))
Walther@60729
  2727
  = (cc, (ist |> path_up_down [R, D] |> upd_env i), body);
Walther@60729
  2728
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2729
val Accept_Tac (Subproblem' _, _, _) =
Walther@60729
  2730
  (*case*) scan_dn cc (ist |> path_down [L, R]) e (*of*);
Walther@60729
  2731
\<close> ML \<open>
Walther@60729
  2732
"~~~~~ fun scan_dn , args:"; val ((cc as (_, ctxt)), (ist as {eval, ...}), t (*fall-through*))
Walther@60729
  2733
  = (cc, (ist |> path_down [L, R]), e);
Walther@60729
  2734
\<close> ML \<open>
Walther@60729
  2735
    (*if*) Tactical.contained_in t (*else*);
Walther@60729
  2736
\<close> ML \<open>
Walther@60729
  2737
val (Program.Tac prog_tac, form_arg) =
Walther@60729
  2738
      (*case*) LItool.check_leaf "next  " ctxt eval (get_subst ist) t (*of*);
Walther@60729
  2739
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2740
           check_tac cc ist (prog_tac, form_arg);
Walther@60729
  2741
\<close> ML \<open>
Walther@60729
  2742
"~~~~~ fun check_tac , args:"; val (((pt, p), ctxt), ist, (prog_tac, form_arg))
Walther@60729
  2743
  = (cc, ist, (prog_tac, form_arg));
Walther@60729
  2744
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2745
    val Subproblem ("Biegelinie", ["normalise", "4x4", "LINEAR", "system"]) =
Walther@60729
  2746
    LItool.tac_from_prog (pt, p) prog_tac;
Walther@60729
  2747
\<close> ML \<open>
Walther@60729
  2748
"~~~~~ fun tac_from_prog , args:"; val ((pt, pos), intac)
Walther@60729
  2749
  = ((pt, p), prog_tac);
Walther@60729
  2750
\<close> ML \<open>
Walther@60729
  2751
    val pos = Pos.back_from_new pos
Walther@60729
  2752
    val ctxt = Ctree.get_ctxt pt pos
Walther@60729
  2753
    val thy = Proof_Context.theory_of ctxt
Walther@60729
  2754
\<close> ML \<open>
Walther@60729
  2755
val (ptac as Const (\<^const_name>\<open>Prog_Tac.SubProblem\<close>, _) $ _ $ _) =
Walther@60729
  2756
    (*case*) intac (*of*)
Walther@60729
  2757
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2758
val Subproblem ("Biegelinie", ["normalise", "4x4", "LINEAR", "system"]) =
Walther@60729
  2759
          fst (Sub_Problem.tac_from_prog (pt, pos) ptac) (*might involve problem refinement etc*)
Walther@60729
  2760
\<close> text \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2761
val (Subproblem _, Subproblem' _) =
Walther@60729
  2762
Sub_Problem.tac_from_prog (pt, pos) ptac;
Walther@60729
  2763
\<close> ML \<open>
Walther@60729
  2764
"~~~~~ fun tac_from_prog , args:"; val ((pt, p), (stac as Const (\<^const_name>\<open>Prog_Tac.SubProblem\<close>, _) $ spec' $ ags'))
Walther@60729
  2765
  = ((pt, pos), ptac);
Walther@60729
  2766
\<close> ML \<open>
Walther@60729
  2767
      val (dI, pI, mI) = Prog_Tac.dest_spec spec'
Walther@60729
  2768
(*+*)val ("Biegelinie", ["LINEAR", "system"], ["no_met"]) = Prog_Tac.dest_spec spec'
Walther@60729
  2769
Walther@60729
  2770
      val thy = common_sub_thy (Know_Store.get_via_last_thy dI, Ctree.root_thy pt);
Walther@60729
  2771
      val init_ctxt = Proof_Context.init_global thy
Walther@60729
  2772
	    val ags = TermC.isalist2list ags';
Walther@60729
  2773
\<close> ML \<open>
Walther@60729
  2774
	      (*if*) mI = ["no_met"] (*then*);
Walther@60729
  2775
            val pors = (M_Match.arguments thy ((#model o (Problem.from_store init_ctxt)) pI) ags): O_Model.T
Walther@60729
  2776
\<close> ML \<open>
Walther@60729
  2777
(*+*)if (pors |> O_Model.to_string @{context}) = "[\n" ^
Walther@60729
  2778
  "(1, [\"1\"], #Given, equalities, [\"[0 = - 1 * c_4 / - 1]\", " ^
Walther@60729
  2779
                                    "\"[0 =\n (- 24 * c_4 + - 24 * L * c_3 + 12 * L \<up> 2 * c_2 + 4 * L \<up> 3 * c +\n  - 1 * L \<up> 4 * q_0) /\n - 24]\", " ^
Walther@60729
  2780
                                    "\"[0 = c_2]\", " ^
Walther@60729
  2781
                                    "\"[0 = (2 * c_2 + 2 * L * c + - 1 * L \<up> 2 * q_0) / 2]" ^
Walther@60729
  2782
                                    "\"]), \n" ^
Walther@60729
  2783
  "(2, [\"1\"], #Given, solveForVars, [\"[c]\", \"[c_2]\", \"[c_3]\", \"[c_4]\"]), \n" ^
Walther@60729
  2784
  "(3, [\"1\"], #Find, solution, [\"ss'''\"])]" then () else error "pors CHANGED";
Walther@60729
  2785
\<close> ML \<open>
Walther@60729
  2786
		        val pI' = Refine.refine_ori' init_ctxt pors pI;
Walther@60729
  2787
\<close> ML \<open>
Walther@60729
  2788
(*+*)val ["normalise", "4x4", "LINEAR", "system"] = pI';
Walther@60729
  2789
\<close> ML \<open> (*ERROR exception Empty raised (line 47 of "./basis/List.sml")*)
Walther@60729
  2790
val return_tac_from_prog_step = (pI', pors (* refinement over models with diff.precond only *), 
Walther@60729
  2791
		          (hd o #solve_mets o Problem.from_store init_ctxt) pI');
Walther@60729
  2792
\<close> ML \<open>
Walther@60729
  2793
                                 (Problem.from_store init_ctxt) pI'; (*.., solve_mets = [], ...*)
Walther@60729
  2794
\<close> ML \<open>
Walther@60729
  2795
val [["EqSystem", "normalise", "4x4"]] =
Walther@60729
  2796
                   (#solve_mets o Problem.from_store init_ctxt) pI';
Walther@60729
  2797
\<close> ML \<open>
Walther@60729
  2798
\<close> ML \<open>
Walther@60729
  2799
\<close> ML \<open>
Walther@60729
  2800
\<close> ML \<open> (*|------------ contine me_Check_Postcond_setzeRandbedingungen ------------------------*)
Walther@60729
  2801
(*|------------------- contine me_Check_Postcond_setzeRandbedingungen ------------------------*)
Walther@60729
  2802
(*\------------------- step into me_Check_Postcond_setzeRandbedingungen --------------------//*)
Walther@60729
  2803
\<close> ML \<open> (*\------------ step into me_Check_Postcond_setzeRandbedingungen --------------------//*)
Walther@60729
  2804
\<close> text \<open>
Walther@60729
  2805
val (p,_,f,nxt,_,pt) = return_Check_Postcond_setzeRandbedingungen;
Walther@60729
  2806
\<close> ML \<open>
Walther@60729
  2807
\<close> ML \<open>
Walther@60729
  2808
\<close> ML \<open>
Walther@60729
  2809
\<close> ML \<open>
Walther@60729
  2810
\<close>
Walther@60729
  2811
Walther@60706
  2812
section \<open>===================================================================================\<close>
Walther@60726
  2813
section \<open>=====  ============================================================================\<close>
Walther@60576
  2814
ML \<open>
Walther@60576
  2815
\<close> ML \<open>
Walther@60576
  2816
\<close> ML \<open>
Walther@60576
  2817
\<close> ML \<open>
wneuper@59472
  2818
\<close>
neuper@52102
  2819
Walther@60576
  2820
section \<open>code for copy & paste ===============================================================\<close>
Walther@60576
  2821
ML \<open>
Walther@60576
  2822
"~~~~~ fun xxx , args:"; val () = ();
Walther@60576
  2823
"~~~~~ and xxx , args:"; val () = ();
Walther@60576
  2824
"~~~~~ from fun xxx \<longrightarrow>fun yyy \<longrightarrow>fun zzz , return:"; val () = ();
Walther@60578
  2825
"~~~~~ continue fun xxx"; val () = ();
Walther@60576
  2826
(*if*) (*then*); (*else*);   (*case*) (*of*);  (*return value*); (*in*) (*end*);
Walther@60576
  2827
"xx"
Walther@60629
  2828
^ "xxx"   (*+*) (*+++*) (*keep for continuing YYYYY*) (*isa*) (*isa2*) (**)
Walther@60576
  2829
\<close> ML \<open> (*//----------- adhoc inserted n ----------------------------------------------------\\*)
Walther@60576
  2830
 (*//----------------- adhoc inserted n ----------------------------------------------------\\*)
Walther@60576
  2831
(*\\------------------ adhoc inserted n ----------------------------------------------------//*)
Walther@60576
  2832
\<close> ML \<open> (*\\----------- adhoc inserted n ----------------------------------------------------//*)
Walther@60576
  2833
Walther@60576
  2834
\<close> ML \<open> (*//----------- step into XXXXX -----------------------------------------------------\\*)
Walther@60576
  2835
(*//------------------ step into XXXXX -----------------------------------------------------\\*)
Walther@60629
  2836
(*keep for continuing YYYYY*)
Walther@60578
  2837
\<close> ML \<open> (*------------- continuing XXXXX ------------------------------------------------------*)
Walther@60578
  2838
(*-------------------- continuing XXXXX ------------------------------------------------------*)
Walther@60578
  2839
(*kept for continuing XXXXX*)
Walther@60576
  2840
(*-------------------- stop step into XXXXX --------------------------------------------------*)
Walther@60576
  2841
\<close> ML \<open> (*------------- stop step into XXXXX --------------------------------------------------*)
Walther@60576
  2842
(*\\------------------ step into XXXXX -----------------------------------------------------//*)
Walther@60576
  2843
\<close> ML \<open> (*\\----------- step into XXXXX -----------------------------------------------------//*)
Walther@60576
  2844
Walther@60578
  2845
(*/------------------- check entry to XXXXX -------------------------------------------------\*)
Walther@60578
  2846
(*\------------------- check entry to XXXXX -------------------------------------------------/*)
Walther@60578
  2847
(*/------------------- check within XXXXX ---------------------------------------------------\*)
Walther@60578
  2848
(*\------------------- check within XXXXX ---------------------------------------------------/*)
Walther@60578
  2849
(*/------------------- check result of XXXXX ------------------------------------------------\*)
Walther@60578
  2850
(*\------------------- check result of XXXXX ------------------------------------------------/*)
Walther@60578
  2851
(* final test ... ----------------------------------------------------------------------------*)
Walther@60578
  2852
Walther@60578
  2853
\<close> ML \<open> (*//----------- inserted hidden code ------------------------------------------------\\*)
Walther@60578
  2854
(*//------------------ inserted hidden code ------------------------------------------------\\*)
Walther@60578
  2855
(*\\------------------ inserted hidden code ------------------------------------------------//*)
Walther@60578
  2856
\<close> ML \<open> (*\\----------- inserted hidden code ------------------------------------------------//*)
Walther@60578
  2857
Walther@60576
  2858
\<close> ML \<open> (*//----------- build new fun XXXXX -------------------------------------------------\\*)
Walther@60576
  2859
(*//------------------ build new fun XXXXX -------------------------------------------------\\*)
Walther@60576
  2860
(*\\------------------ build new fun XXXXX -------------------------------------------------//*)
Walther@60576
  2861
\<close> ML \<open> (*\\----------- build new fun XXXXX -------------------------------------------------//*)
Walther@60710
  2862
\<close> ML \<open>
Walther@60576
  2863
\<close>
neuper@52102
  2864
end