test/Tools/isac/MathEngine/mathengine-stateless.sml
author wneuper <walther.neuper@jku.at>
Tue, 20 Apr 2021 16:58:44 +0200
changeset 60242 73ee61385493
parent 60230 0ca0f9363ad3
child 60255 5497a3d67d96
permissions -rw-r--r--
replace power ^^^ by \<up>
     1 (* Title:  tests for MathEngine/mathengine-stateless.sml
     2    Author: Walther Neuper 2000, 2006, 2020
     3    (c) due to copyright terms
     4 
     5 theory Test_Some imports Build_Thydata begin 
     6 ML {* KEStore_Elems.set_ref_thy @{theory};
     7   fun autoCalculate' cI auto = autoCalculate cI auto |> Future.join *}
     8 ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Interpret/mathengine.sml"
     9 *)
    10 "--------------------------------------------------------";
    11 "table of contents --------------------------------------";
    12 "--------------------------------------------------------";
    13 "----------- change to TermC.parse ctxt -----------------------";
    14 "----------- debugging setContext..pbl_ -----------------";
    15 "----------- tryrefine ----------------------------------";
    16 "----------- search for Or_to_List ----------------------";
    17 "----------- check thy in CalcTreeTEST ------------------";
    18 "----------- identified error in fun getTactic, ThmC.string_of_thm ---------------";
    19 "----------- improved fun getTactic for interSteps and numeral calculations --";
    20 "--------------------------------------------------------";
    21 "--------------------------------------------------------";
    22 "--------------------------------------------------------";
    23 
    24 "----------- change to TermC.parse ctxt -----------------------";
    25 "----------- change to TermC.parse ctxt -----------------------";
    26 "----------- change to TermC.parse ctxt -----------------------";
    27 "===== start calculation: from problem description (fmz) to origin";
    28 val fmz = ["realTestGiven (((1+2)*4/3) \<up> 2)", "realTestFind s"];
    29 val (thyID, pblID, metID) =
    30   ("Test", ["calculate", "test"], ["Test", "test_calculate"]);
    31 (*======= Isabelle2013-2 --> Isabelle2014: unclear, why this test ever run =====================*)
    32 
    33 "----------- debugging setContext..pbl_ -----------------";
    34 "----------- debugging setContext..pbl_ -----------------";
    35 "----------- debugging setContext..pbl_ -----------------";
    36 reset_states ();
    37 CalcTree [(["equality (x+1=(2::real))", "solveFor x", "solutions L"], 
    38   ("Test", ["sqroot-test", "univariate", "equation", "test"],
    39    ["Test", "squ-equ-test-subpbl1"]))];
    40 Iterator 1;
    41 moveActiveRoot 1; modelProblem 1;
    42 
    43 val pos as (p,_) = ([],Pbl);
    44 val guh = "pbl_equ_univ";
    45 (********************************)checkContext 1 pos guh;(**************************************)
    46 val ((pt, ([], Pbl)), []) = get_calc 1;
    47 val pp = par_pblobj pt p;
    48 val keID = Ptool.guh2kestoreID guh;
    49 case context_pbl keID pt pp of (true, ["univariate", "equation"], _, _, _) => ()
    50 | _ => error "mathengine.sml: context_pbl .. pbl_equ_univ checked";
    51 
    52 case get_obj g_spec pt p of (_, ["empty_probl_id"], _) => ()
    53 | _ => error "mathengine.sml: context_pbl .. pbl still empty";
    54 (********************************)setContext 1 pos guh;(***************************************)
    55 val ((pt, ([], Pbl)), []) = get_calc 1;
    56 case get_obj g_spec pt p of (_, ["univariate", "equation"], _) => ()
    57 | _ => error "mathengine.sml: context_pbl .. pbl set";
    58 
    59 
    60 setContext 1 pos "met_eq_lin";
    61 val ((pt,_),_) = get_calc 1;
    62 case get_obj g_spec pt p of
    63   ("empty_thy_id", ["univariate", "equation"], ["LinEq", "solve_lineq_equation"]) => ()
    64 | _ => error "mathengine.sml: context_pbl .. pbl set";
    65 
    66 
    67 "----------- tryrefine ----------------------------------";
    68 "----------- tryrefine ----------------------------------";
    69 "----------- tryrefine ----------------------------------";
    70 reset_states ();
    71 CalcTree [(["equality (x/(x^2 - 6*x+9) - 1/(x^2 - 3*x) =1/x)", 
    72 	    "solveFor x", "solutions L"],
    73 	   ("RatEq",["univariate", "equation"],["no_met"]))];
    74 Iterator 1;
    75 moveActiveRoot 1; autoCalculate 1 CompleteCalc;
    76 
    77 refineProblem 1 ([1],Res) "pbl_equ_univ" 
    78 (*gives "pbl_equ_univ_rat" correct*);
    79 
    80 refineProblem 1 ([1],Res) (Ptool.pblID2guh ["LINEAR", "univariate", "equation"])
    81 (*gives "pbl_equ_univ_lin" incorrect*);
    82 
    83 "--------- search for Or_to_List ------------------------";
    84 "--------- search for Or_to_List ------------------------";
    85 "--------- search for Or_to_List ------------------------";
    86 val fmz = ["equality (x \<up> 2 + 4*x + 5 = (2::real))", "solveFor x", "solutions L"];
    87 val (dI',pI',mI') = ("Isac_Knowledge", ["univariate", "equation"], ["no_met"])
    88 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
    89 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
    90 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
    91 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
    92 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
    93 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
    94 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
    95 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
    96 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
    97 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
    98 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = ("Subproblem"*)
    99 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   100 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   101 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   102 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   103 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   104 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   105 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   106 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   107 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   108 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   109 "~~~~~ fun me, args:"; val (tac, (p:pos'), (_:NEW(*remove*)), (pt:ctree)) = 
   110                             (nxt, p, [], pt);
   111 val ("ok", (_, _, ptp))  = Step.by_tactic tac (pt,p)
   112 val (pt, p) = ptp;
   113 "~~~~~ fun Step.do_next, args:"; val (((ip as (_,p_)):pos'), ((ptp as (pt,p), tacis):Calc.state_pre)) =
   114                           (p, ((pt, e_pos'),[]));
   115 val pIopt = get_pblID (pt,ip);
   116 ip = ([],Res); (*false*)
   117 ip = p; (*false*)
   118 member op = [Pbl,Met] p_; (*false*)
   119 "~~~~~ fun do_next, args:"; val (ptp as (pt, pos as (p,p_))) = (pt,ip);
   120 MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p); (*false*)
   121 val thy' = get_obj g_domID pt (par_pblobj pt p);
   122 val ((ist, ctxt), sc) = resume_prog thy' (p,p_) pt;
   123 
   124 val Next_Step (istate, ctxt, tac) = LI.find_next_step sc (pt,pos) ist ctxt; (*WAS Empty_Tac_: tac_*)
   125 case tac of Or_to_List' _ => ()
   126 | _ => error "Or_to_List broken ?"
   127 
   128 
   129 "----------- check thy in CalcTreeTEST ------------------";
   130 "----------- check thy in CalcTreeTEST ------------------";
   131 "----------- check thy in CalcTreeTEST ------------------";
   132 "WN1109 with Inverse_Z_Transform.thy when starting tests with CalcTreeTEST \n" ^
   133 "we got the message: ME_Isa: thy 'Build_Inverse_Z_Transform' not in system \n" ^
   134 "Below there are the steps which found out the reason: \n" ^
   135 "store_pbt mistakenly stored that theory.";
   136 val ctxt = Proof_Context.init_global @{theory Isac_Knowledge};
   137 val SOME t = parseNEW ctxt "filterExpression (X  = 3 / (z - 1/4 + -1/8 * (1/(z::real))))";
   138 val SOME t = parseNEW ctxt "stepResponse (x[n::real]::bool)";
   139 
   140 val fmz = ["filterExpression (X  = 3 / (z - 1/4 + -1/8 * (1/(z::real))))", "boundVariable z",
   141   "stepResponse (x[n::real]::bool)"];
   142 val (dI,pI,mI) = ("Isac_Knowledge", ["Inverse", "Z_Transform", "SignalProcessing"], 
   143   ["SignalProcessing", "Z_Transform", "Inverse"]);
   144 
   145 val (p,_,fb,nxt,_,pt)  = CalcTreeTEST [(fmz, (dI,pI,mI))]; 
   146 (*WAS ME_Isa: thy 'Build_Inverse_Z_Transform' not in system*)
   147 
   148 "~~~~~ T fun CalcTreeTEST, args:"; val [(fmz, sp):Formalise.T] = [(fmz, (dI,pI,mI))];
   149 "~~~~~ fun nxt_specify_init_calc, args:"; val (fmz, (dI, pI, mI)) = (fmz, sp);
   150 	      val thy = ThyC.get_theory dI;
   151     mI = ["no_met"]; (*false*)
   152 	          val pors = Problem.from_store pI |> #ppc |> O_Model.init fmz thy;
   153             val pctxt = ContextC.initialise' thy fmz;
   154 	      val {cas, ppc, thy=thy',...} = Problem.from_store pI; (*take dI from _refined_ pbl*)
   155 	"tracing bottom up"; Context.theory_name thy' = "Build_Inverse_Z_Transform"; (*WAS true*)
   156       val dI = Context.theory_name (ThyC.parent_of thy thy');
   157 "tracing bottom up"; dI = "Build_Inverse_Z_Transform"; (*WAS true*)
   158     cas = NONE; (*true*)
   159 	      val hdl = pblterm dI pI;
   160         val (pt, _) = cappend_problem e_ctree [] (Istate.empty, pctxt) (fmz, (dI, pI, mI))
   161 				  (pors, (dI, pI, mI), hdl, ContextC.empty)
   162 ;
   163 "~~~~~ fun Step_Specify.by_tactic_input, args:"; val ((tac as Model_Problem), (pt, pos as (p,p_))) = (Model_Problem, (pt, ([], Pbl)));
   164       val PblObj {origin = (oris, ospec, _), probl, spec, ctxt, ...} = get_obj I pt p;
   165 "tracing bottom up"; #1(ospec) = "Build_Inverse_Z_Transform"; (*WAS true*)
   166 
   167 "~~~~~ fun References.select, args:"; val ((odI, opI, omI), (dI, pI, mI)) = (ospec, spec);
   168 dI = ThyC.id_empty; (*true*)
   169 odI = "Build_Inverse_Z_Transform"; (*true*)
   170 dI = "empty_thy_id"; (*true*)
   171 "~~~~~ after fun References.select:";
   172       val (dI, pI, mI) = References.select ospec spec;
   173 "tracing bottom up"; dI = "Build_Inverse_Z_Transform"; (*WAS true*)
   174       val thy = ThyC.get_theory dI; (*caused ME_Isa: thy 'Build_Inverse_Z_Transform' not in system*)
   175 
   176 "----------- identified error in fun getTactic, ThmC.string_of_thm ---------------";
   177 "----------- identified error in fun getTactic, ThmC.string_of_thm ---------------";
   178 "----------- identified error in fun getTactic, ThmC.string_of_thm ---------------";
   179 reset_states ();
   180 CalcTree [(["equality (x+1=(2::real))", "solveFor x", "solutions L"],
   181   ("Test",["sqroot-test", "univariate", "equation", "test"],["Test", "squ-equ-test-subpbl1"]))];
   182 moveActiveRoot 1;
   183 autoCalculate 1 CompleteCalcHead;
   184 autoCalculate 1 (Steps 1);
   185 autoCalculate 1 (Steps 1);
   186 autoCalculate 1 (Steps 1);
   187 getTactic 1 ([1],Frm); (*<RULESET> norm_equation </RULESET>*)
   188 interSteps 1 ([1],Res);
   189 "~~~~~ fun interSteps, args:"; val (cI, ip) = (1, ([1],Res));
   190 val ((pt,p), tacis) = get_calc cI;
   191 (not o is_interpos) ip = false;
   192 val ip' = lev_pred' pt ip;
   193 "~~~~~ fun go, args:"; val (pt, (pos as (p,p_):pos')) = (pt, ip);
   194 (*         \<up> \<up> \<up> not in test/../ *)
   195     val nd = get_nd pt p
   196     val cn = children nd;
   197 null cn = false;
   198 (is_rewset o (get_obj g_tac nd)) [(*root of nd*)] = true;
   199 "~~~~~ fun detailrls, args:"; val (pt, (pos as (p,p_):pos')) = (pt, pos);
   200 (*         \<up> \<up> \<up> only once in test/../solve.sml*)
   201     val t = get_obj g_form pt p
   202 	  val tac = get_obj g_tac pt p
   203 	  val rls = (assoc_rls o rls_of) tac
   204     val ctxt = get_ctxt pt pos
   205 (*rls = Rule_Set.Repeat {calc = [], erls = ...*)
   206           val is = Istate.init_detail tac t
   207 	        (*TODO.WN060602 Pstate (["(t_, Problem (Isac,[equation,univar]))"]
   208 				    is wrong for simpl, but working ?!? *)
   209 	        val tac_ = Apply_Method' (MethodC.id_empty(*WN0402: see Step.add !?!*), SOME t, is, ctxt)
   210 	        val pos' = ((lev_on o lev_dn) p, Frm)
   211 	        val thy = ThyC.get_theory "Isac_Knowledge"
   212 	        val (_,_,_,pt') = (*implicit Take*)Step.add tac_ (is, ctxt) (pt, pos');
   213 	        (*val (_,_,(pt'',_)) = *)complete_solve CompleteSubpbl [] (pt',pos');
   214 	        (*                   \<up> \<up> \<up> \<up> ^^ in test/../mathengine.sml*)
   215 	        (* in pt'': tac = Rewrite ("rnorm_equation_add", "Test.rnorm_equation_add")},
   216                                                            \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up>  \<up> *)
   217 "~~~~~ fun complete_solve , args:"; val (auto, c, (ptp as (_, p as (_,p_)): ctree * pos')) =
   218   (CompleteSubpbl, [], (pt',pos'));
   219 p = ([], Res) = false;
   220 member op = [Pbl,Met] p_ = false;
   221 val (_, (_, c', ptp')) = Step_Solve.do_next ptp;
   222 (* in pt': tac = Rewrite ("rnorm_equation_add", "Test.rnorm_equation_add")},
   223                                                 \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up>  \<up> *)
   224 "~~~~~ fun do_next , args:"; val ((ptp as (pt, pos as (p,p_)))) = (ptp);
   225 MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p) = false;
   226           val thy' = get_obj g_domID pt (par_pblobj pt p);
   227 	        val ((ist, ctxt), sc) = resume_prog thy' (p,p_) pt;
   228 	        val Next_Step (is, ctxt, tac_) = LI.find_next_step sc (pt,pos) ist ctxt;
   229 (*tac_ = Rewrite' (..., ("rnorm_equation_add", "Test.rnorm_equation_add"), ...) !!!!!!!!!!!!!!!!*)
   230 
   231 (*----- outcommented during cleanup of args in lucas-interpreter.sml ------------------------\\* )
   232 "~~~~~ fun find_next_step , args:"; val (thy, (ptp as (pt, pos as (p, _)):ctree * pos'), 
   233   (sc as Prog (h $ body)), (Pstate (ist as (E,l,rls,a,v,s,b)), ctxt)) = ((thy',srls), (pt,pos), sc, is);
   234 l = [] = false;
   235 go_scan_up thy ptp sc ist true (*--> Accept_Tac (Rewrite' (.., (.., "Test.rnorm_equation_add"), ...)
   236   BUT WE FIND IN THE CODE ABOVE IN find_next_step:*)
   237 ;
   238 (*----------------------------------------------------------------------------------------------*)
   239 if ThmC.string_of_thm @{thm rnorm_equation_add} =  "\<not> ?b =!= 0 \<Longrightarrow> (?a = ?b) = (?a + - 1 * ?b = 0)"
   240 then () else error "ThmC.string_of_thm changed";
   241 if ThmC.string_of_thm @{thm rnorm_equation_add} =  "\<not> ?b =!= 0 \<Longrightarrow> (?a = ?b) = (?a + - 1 * ?b = 0)"
   242 then () else error "string_of_thm changed";
   243 (*----------------------------------------------------------------------------------------------*)
   244 ;
   245 (*SEARCH FOR THE ERROR WENT DOWN TO THE END OF THIS TEST, AND THEN UP TO HERE AGAIN*)
   246 "~~~~~ fun begin_end_prog, args:"; val (tac_, is, (pt, pos as (p,p_))) = (tac_, is, ptp);
   247         val pos =
   248           case pos of
   249 		        (p, Met) => ((lev_on o lev_dn) p, Frm)(*begin script*)
   250 		      | (p, Res) => (lev_on p,Res) (*somewhere in script*)
   251 		      | _ => pos;
   252 Step.add tac_ is pos pt;
   253 (* tac = Rewrite ("rnorm_equation_add", "Test.rnorm_equation_add")},
   254                                         \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up>  \<up> *)
   255 "~~~~~ fun add, args:"; val (thy, (Rewrite' (thy',ord',rls',pa,thm',f,(f',asm))), 
   256   (is, ctxt), (p,p_), pt) = ((ThyC.get_theory "Isac_Knowledge"), tac_, is, pos, pt);
   257         val (pt,c) = cappend_atomic pt p (is, ContextC.insert_assumptions asm ctxt) f
   258           (Rewrite thm') (f',asm) Complete;
   259 (* in pt: tac = Rewrite ("rnorm_equation_add", "Test.rnorm_equation_add")},
   260                                                \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up>  \<up> *)
   261 "~~~~~ fun cappend_atomic, args:"; val (pt, p, ist_res, f, r, f', s) = 
   262   (pt, p, (is, ContextC.insert_assumptions asm ctxt), f, (Rewrite thm'), (f',asm), Complete);
   263 existpt p pt andalso Tactic.is_empty (get_obj g_tac pt p) = false;
   264 apfst (append_atomic p ist_res f r f' s) (cut_tree pt (p,Frm));
   265 apfst : ('a -> 'b) -> 'a * 'c -> 'b * 'c;
   266 (append_atomic p ist_res f r f' s) : ctree -> ctree;
   267 ;
   268 (* HERE THE ERROR OCCURED IN THE FIRST APPROACH
   269   getTactic 1 ([1,1],Frm); <ERROR> syserror in getTactic </ERROR>  <<<<<=========================*)
   270 "~~~~~ fun getTactic, args:"; val (cI, (p:pos')) = (1, ([1,1],Frm));
   271 val ((pt,_),_) = get_calc cI
   272 val (form, tac, asms) = ME_Misc.pt_extract (pt, p)
   273 val SOME ta = tac;
   274 "~~~~~ fun gettacticOK2xml, args:"; val ((cI:calcID), tac) = (cI, ta);
   275 val i = 2;
   276 "~~~~~ fun tac2xml, args:"; val (j, (Rewrite thm')) = (i, tac);
   277 "~~~~~ fun thm'2xml, args:"; val (j, ((ID, form) : thm'')) = ((j+i), thm');
   278 ID = "rnorm_equation_add";
   279 @{thm rnorm_equation_add};
   280 (UnparseC.term o Thm.prop_of) form = "~ ?b =!= 0 ==> (?a = ?b) = (?a + -1 * ?b = 0)"
   281   (*?!? should be "\<not> ?b =!= 0 \<Longrightarrow> (?a = ?b) = (?a + -1 * ?b = 0)"*)
   282 (*thmstr2xml (j+i) form;
   283 ERROR Undeclared constant: "Test.rnorm_equation_add" \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up> \<up> ^*)
   284 ;
   285 Test_Tool.show_pt pt;
   286 (*[
   287 (([], Frm), solve (x + 1 = 2, x)),
   288 (([1], Frm), x + 1 = 2),
   289 (([1,1], Frm), x + 1 = 2),
   290 (([1,1], Res), x + 1 + -1 * 2 = 0),
   291 (([1], Res), x + 1 + -1 * 2 = 0),
   292 (([2], Res), -1 + x = 0)]
   293 
   294 pt;   --> tac = Rewrite ("rnorm_equation_add", "Test.rnorm_equation_add")*)
   295 ( *----- outcommented during cleanup of args in lucas-interpreter.sml ------------------------//*)
   296 
   297 "----------- improved fun getTactic for interSteps and numeral calculations --";
   298 "----------- improved fun getTactic for interSteps and numeral calculations --";
   299 "----------- improved fun getTactic for interSteps and numeral calculations --";
   300 reset_states (); val cI = 1;
   301 CalcTree [(["equality (x+1=(2::real))", "solveFor x", "solutions L"],
   302   ("Test",["sqroot-test", "univariate", "equation", "test"],["Test", "squ-equ-test-subpbl1"]))];
   303 moveActiveRoot 1;
   304 autoCalculate 1 CompleteCalcHead;
   305 autoCalculate 1 (Steps 1);
   306 
   307     val cs = get_calc cI             (* we improve from the calcstate here [*] *);
   308     val pos as (_,p_) = get_pos cI 1 (* we improve from the calcstate here [*] *);
   309 
   310 appendFormula 1 "x + 4 + -3 = 2" (*|> Future.join*);
   311 interSteps 1 ([1],Res); (* already done by appendFormula, now showed to frontend *)
   312 getTactic 1 ([1,1], Res); 
   313   (*<ID> sym_#add_Float ((~3,0), (0,0)) __ ((4,0), (0,0)) </ID>  <<<<<================== to improve
   314   <ISA> 1 + x = -3 + (4 + x) </ISA>*)
   315 
   316 val ((pt''',p'''), tacis''') = get_calc cI;
   317 Test_Tool.show_pt pt'''
   318 (*[
   319 (([], Frm), solve (x + 1 = 2, x)),
   320 (([1], Frm), x + 1 = 2),
   321 (([1,1], Frm), x + 1 = 2),
   322 (([1,1], Res), 1 + x = 2),
   323 (([1,2], Res), -3 + (4 + x) = 2),
   324 (([1,3], Res), -3 + (x + 4) = 2),
   325 (([1,4], Res), x + 4 + -3 = 2),
   326 (([1], Res), x + 4 + -3 = 2)]*)
   327 ;
   328 "~~~~~ fun appendFormula', args:"; val (cI, ifo) = (1, "x + 4 + -3 = 2");
   329 (*  val cs = get_calc cI             (* we improve from the calcstate here [*] *);
   330     val pos as (_,p_) = get_pos cI 1 (* we improve from the calcstate here [*] *);*)
   331 val ("ok", cs' as (_, _, ptp')) = Step.do_next pos cs;
   332 (*val ("ok", (_, c, ptp as (_,p))) = *)Step_Solve.by_term ptp' (encode ifo);
   333 "~~~~~ fun Step_Solve.by_term , args:"; val (cs as (_, _, ptp as (pt, pos as (p, p_))), istr) =
   334   (cs', (encode ifo));
   335 val SOME f_in = TermC.parse (ThyC.get_theory "Isac_Knowledge") istr
   336 	      val f_in = Thm.term_of f_in
   337 	      val f_succ = get_curr_formula (pt, pos);
   338 f_succ = f_in  = false;
   339 val NONE = CAS_Cmd.input f_in 
   340 			          val pos_pred = lev_back' pos
   341 			          (* f_pred ---"step pos cs"---> f_succ in appendFormula *)
   342 			          val f_pred = get_curr_formula (pt, pos_pred);
   343 			          (*val msg_calcstate' = *)compare_step ([], [], (pt, pos_pred)) f_in (*<<=====*);
   344 "~~~~~ fun compare_step, args:"; val (((tacis, c, ptp as (pt, pos as (p,p_)))), ifo) = 
   345   (([], [], (pt, pos_pred)), f_in);
   346     val fo =
   347       case p_ of
   348         Frm => get_obj g_form pt p
   349 			| Res => (fst o (get_obj g_result pt)) p
   350 			| _ => TermC.empty (*on PblObj is fo <> ifo*);
   351 	  val {nrls, ...} = MethodC.from_store (get_obj g_metID pt (par_pblobj pt p))
   352 	  val {rew_ord, erls, rules, ...} = Rule_Set.rep nrls;
   353 	  (*val (found, der) = *)Derive.steps rew_ord erls rules fo ifo; (*<---------------*)
   354 "~~~~~ fun .concat_deriv, args:"; val (rew_ord, erls, rules, fo, ifo) =
   355   (rew_ord, erls, rules, fo, ifo);
   356     fun derivat ([]:(term * rule * (term * term list)) list) = TermC.empty
   357       | derivat dt = (#1 o #3 o last_elem) dt
   358     fun equal (_,_,(t1, _)) (_,_,(t2, _)) = t1=t2
   359     (*val  fod = *)Derive.do_one (Isac"") erls rules (snd rew_ord) NONE  fo;
   360     (*val ifod = *)Derive.do_one (Isac"") erls rules (snd rew_ord) NONE ifo;
   361