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