test/Tools/isac/Test_Isac_Short.thy
author wneuper <walther.neuper@jku.at>
Thu, 15 Jul 2021 14:10:18 +0200
changeset 60324 5c7128feb370
parent 60318 e6e7a9b9ced7
child 60325 a7c0e6ab4883
permissions -rw-r--r--
ewrite.sml + poly.sml + rational.sml: ok, repair rewrite-orders
     1 (* Title:  All tests on isac (some outcommented since Isabelle2002-->2009-2)
     2    Author: Walther Neuper 101001
     3    (c) copyright due to license terms.
     4 
     5    Isac's tests are organised parallel to sources: 
     6      $ISABELLE_ISAC_TEST has same directory structure as $ISABELLE_ISAC
     7    plus
     8      $ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS
     9      $ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl: the Lucas-Interpreter's core functionality
    10 
    11 Note, that only the first error in a file is shown here.
    12 *)
    13 
    14 section \<open>Notes on running tests\<close>
    15 subsection \<open>Switch between running tests and updating code\<close>
    16 text \<open>
    17   Isac encapsulates code as much as possible in structures without open.
    18   This policy conflicts with those tests, which go into functions to details
    19   not declared in the signatures.
    20 \<close>
    21 subsection \<open>Decide between running Test_Isac_Short.thy and Test_Isac.thy\<close>
    22 text \<open>
    23   Some tests raise exception Size raised (line 171 of "./basis/LibrarySupport.sml")
    24   if run in x86_64_32 mode of Poly/ML 5.8 (which is set as default).
    25   This exception can be avoided by ML_system_64 = "true" in ~/.isabelle/isabisac/etc/preferences.
    26   A model is in the repository at ~~/etc/preferences.
    27   These preferences have drawbacks, however:
    28   * they claim more memory such that Isabelle instances canNOT run in parallel.
    29   * they do NOT reach Build_Isac.thy hanging in Build_Thydata.thy, see there.
    30 
    31   So default for Build_Isac.thy and for general testing is Test_Isac_Short.thy is x86_64_32 mode.
    32   From time to time full testing in Test_Isac.thy is recommended. For that purpose
    33   * set ML_system_64 = "true" in ~/.isabelle/isabisac/etc/preferences.
    34 
    35 \\******************* don't forget to re-set defaults BEFORE updating code *******************//
    36 
    37     Note that Isabelle/jEdit re-generates the preferences file on shutdown, thus always use
    38     ***************** $ gedit ~/.isabelle/isabisac/etc/preferences &
    39 \<close>
    40 
    41 section \<open>Run the tests\<close>
    42 text \<open>
    43 * say "OK" to the popup asking for theories to be loaded
    44 * watch the <Theories> window for errors in the "imports" below
    45 \<close>
    46 
    47 theory Test_Isac_Short
    48   imports Isac.Build_Isac (* note that imports are WITHOUT open struct ..*)
    49   (* in case of ERROR Bad theory import "Draft.Thy_All"..., open each theory one by one
    50      and find out, which ML_file or *.thy causes an error (might be ONLY one).
    51      Also backup files (#* ) recognised by jEdit cause this trouble                    *)
    52 (*/---------------------- do Minisubpbl before ADDTESTS/All_Ctxt ------------------------------\*)
    53 (** )"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/accumulate-val/Thy_All"( *TODOO*)
    54 (**)"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/Ctxt"
    55 (**)"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/test-depend/Build_Test"
    56 (**)"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/All_Ctxt"
    57 (**)"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/Test_Units"
    58 (**)"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/course/phst11/T1_Basics"
    59 (** )"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/course/phst11/T2_Rewriting"( *TODOO*)
    60 (**)"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/course/phst11/T3_MathEngine"
    61 (**)"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/file-depend/BuildC_Test"
    62 (**)"$ISABELLE_ISAC_TEST/Tools/isac/ADDTESTS/session-get_theory/Foo"
    63 (*"ADDTESTS/course/SignalProcess/Build_Inverse_Z_Transform"
    64    ADDTESTS/------------------------------------------- see end of tests *)
    65 (*/~~~ these work directly from Pure, but create problems here ..
    66   "$ISABELLE_ISAC_TEST/Pure/Isar/Keyword_ISAC.thy"           (* Malformed theory import, "keywords" ?!? *)
    67   "$ISABELLE_ISAC_TEST/Pure/Isar/Test_Parse_Isac.thy"        (* Malformed theory import, "keywords" ?!? *)
    68   "$ISABELLE_ISAC_TEST/Pure/Isar/Test_Parsers_Cookbook.thy"  (* Malformed theory import             ?!? *)
    69   "$ISABELLE_ISAC_TEST/Pure/Isar/Theory_Commands"            (* Duplicate outer syntax command "ISAC"   *)
    70   "$ISABELLE_ISAC_TEST/Pure/Isar/Downto_Synchronized"        (* re-defines / breaks structures      !!! *)
    71   \~~~ .. these work independently, but create problems here *)
    72 (**)"$ISABELLE_ISAC_TEST/Pure/Isar/Test_Parsers"
    73 (**)"$ISABELLE_ISAC_TEST/HOL/Tools/Sledgehammer/Try_Sledgehammer"
    74 (*/---------------------- do Minisubpbl before ADDTESTS/All_Ctxt ------------------------------\*)
    75   "$ISABELLE_ISAC_TEST/Tools/isac/Specify/refine"        (* setup for refine.sml   *)
    76   "$ISABELLE_ISAC_TEST/Tools/isac/ProgLang/calculate"    (* setup for evaluate.sml *)
    77   "$ISABELLE_ISAC_TEST/Tools/isac/Knowledge/integrate"   (* setup for integrate.sml*)
    78 (** )
    79 (*\---------------------- do Minisubpbl before ADDTESTS/All_Ctxt ------------------------------/*)
    80 (*"$ISABELLE_ISAC/Knowledge/GCD_Poly_OLD" (*not imported by Isac.thy*)        Test_Isac_Short*)
    81 (*"$ISABELLE_ISAC/Knowledge/GCD_Poly_FP"  (*not imported by Isac.thy*)        Test_Isac_Short*)
    82 (*\---------------------- do Minisubpbl before ADDTESTS/All_Ctxt ------------------------------/*)
    83 ( **)
    84 
    85 begin
    86 
    87 declare [[ML_print_depth = 20]]
    88 
    89 ML \<open>open ML_System\<close>
    90 ML \<open>
    91   open Kernel;
    92   open Math_Engine;
    93   open Test_Code;              CalcTreeTEST;
    94   open LItool;                 arguments_from_model;
    95   open Sub_Problem;
    96   open Fetch_Tacs;
    97   open Step
    98   open Env;
    99   open LI;                     scan_dn;
   100   open Istate;
   101   open Error_Pattern;
   102   open Error_Pattern_Def;
   103   open Specification;
   104   open Ctree;                  append_problem;
   105   open Pos;
   106   open Program;
   107   open Prog_Tac;
   108   open Tactical;
   109   open Prog_Expr;
   110   open Auto_Prog;              rule2stac;
   111   open Input_Descript;
   112   open Specify;
   113   open Specify;
   114   open Step_Specify;
   115   open Step_Solve;
   116   open Step;
   117   open Solve;                  (* NONE *)
   118   open ContextC;               transfer_asms_from_to;
   119   open Tactic;                 (* NONE *)
   120   open I_Model;
   121   open O_Model;
   122   open P_Model;                (* NONE *)
   123   open Rewrite;
   124   open Eval;                   get_pair;
   125   open TermC;                  atomt;
   126   open Rule;
   127   open Rule_Set;               Sequence;
   128   open Eval_Def
   129   open ThyC
   130   open ThmC_Def
   131   open ThmC
   132   open Rewrite_Ord
   133   open UnparseC
   134 \<close>
   135 
   136 ML \<open>
   137 "~~~~~ fun xxx , args:"; val () = ();
   138 "~~~~~ and xxx , args:"; val () = ();
   139 "~~~~~ from fun xxx \<longrightarrow>fun yyy \<longrightarrow>fun zzz , return:"; val () = ();
   140 (*if*) (*then*); (*else*);   (*case*) (*of*);  (*return value*); (*in*) (*end*);
   141 "xx"
   142 ^ "xxx"   (*+*) (*!for return!*) (*isa*) (*REP*) (**)
   143 \<close> ML \<open>
   144 \<close>
   145 ML \<open>
   146 \<close> ML \<open>
   147 \<close> ML \<open>
   148 \<close> ML \<open>
   149 \<close> ML \<open>
   150 \<close> ML \<open>
   151 \<close> ML \<open>
   152 \<close> ML \<open>
   153 \<close> ML \<open>
   154 \<close> ML \<open>
   155 \<close> ML \<open>
   156 \<close>
   157 
   158 ML \<open>
   159   KEStore_Elems.set_ref_thy @{theory};
   160   (*fun autoCalculate' cI auto = autoCalculate cI auto (*|> Future.join*)*);
   161 \<close>
   162 
   163 (*---------------------- check test file by testfile -------------------------------------------
   164   ---------------------- check test file by testfile -------------------------------------------*)
   165 section \<open>trials with Isabelle's functions\<close>
   166   ML \<open>"%%%%%%%%%%%%%%%%% start Isabelle %%%%%%%%%%%%%%%%%%%%%%%";\<close>
   167   ML_file "$ISABELLE_ISAC_TEST/Pure/General/alist.ML"
   168   ML_file "$ISABELLE_ISAC_TEST/Pure/General/basics.ML"
   169   ML_file "$ISABELLE_ISAC_TEST/Pure/General/scan.ML"
   170   ML_file "$ISABELLE_ISAC_TEST/Pure/PIDE/xml.ML"
   171   ML \<open>"%%%%%%%%%%%%%%%%% end Isabelle %%%%%%%%%%%%%%%%%%%%%%%%%";\<close>
   172 
   173 section \<open>test ML Code of isac\<close>
   174 subsection \<open>basic code first\<close>
   175   ML \<open>"%%%%%%%%%%%%%%%%% start ProgLang.thy %%%%%%%%%%%%%%%%%%%";\<close>
   176   ML_file "BaseDefinitions/libraryC.sml"
   177   ML_file "BaseDefinitions/rule-def.sml"
   178   ML_file "BaseDefinitions/eval-def.sml"
   179   ML_file "BaseDefinitions/rewrite-order.sml"
   180   ML_file "BaseDefinitions/theoryC.sml"
   181   ML_file "BaseDefinitions/rule.sml"
   182   ML_file "BaseDefinitions/thmC-def.sml"
   183   ML_file "BaseDefinitions/error-fill-def.sml"
   184   ML_file "BaseDefinitions/rule-set.sml"
   185   ML_file "BaseDefinitions/check-unique.sml"
   186 (*called by Know_Store..*)
   187   ML_file "BaseDefinitions/calcelems.sml"
   188   ML_file "BaseDefinitions/termC.sml"
   189   ML_file "BaseDefinitions/substitution.sml"
   190 (** )ML_file "BaseDefinitions/contextC.sml"( *loops with eliminate ThmC.numerals_to_Free*)
   191   ML_file "BaseDefinitions/environment.sml"
   192 (** )ML_file "BaseDefinitions/kestore.sml" ( * setup in ADDTEST/accumulate-val/lucas_interpreter.sml*)
   193 (*---------------------- do Minisubpbl before ADDTESTS/All_Ctxt --------------------------------
   194   ---------------------- do Minisubpbl before ADDTESTS/All_Ctxt --------------------------------*)
   195 
   196   ML_file "ProgLang/calculate.sml"
   197   ML_file "ProgLang/evaluate.sml"       (* requires setup from calculate.thy                    *)
   198   ML_file "ProgLang/listC.sml"
   199   ML_file "ProgLang/prog_expr.sml"
   200   ML_file "ProgLang/program.sml"
   201   ML_file "ProgLang/prog_tac.sml"
   202   ML_file "ProgLang/tactical.sml"
   203 (** )ML_file "ProgLang/auto_prog.sml"( *loops with eliminate ThmC.numerals_to_Free*)
   204 (*---------------------- do Minisubpbl before ADDTESTS/All_Ctxt --------------------------------
   205   ---------------------- do Minisubpbl before ADDTESTS/All_Ctxt --------------------------------*)
   206 
   207 subsection \<open>basic functionality on simple example first\<close>
   208   ML_file "Minisubpbl/000-comments.sml"
   209   ML_file "Minisubpbl/100-init-rootpbl.sml"
   210   ML_file "Minisubpbl/150-add-given.sml"
   211   ML_file "Minisubpbl/200-start-method-NEXT_STEP.sml"
   212   ML_file "Minisubpbl/200-start-method.sml"
   213   ML_file "Minisubpbl/250-Rewrite_Set-from-method.sml"
   214   ML_file "Minisubpbl/300-init-subpbl-NEXT_STEP.sml"
   215   ML_file "Minisubpbl/300-init-subpbl.sml"
   216   ML_file "Minisubpbl/400-start-meth-subpbl.sml"
   217   ML_file "Minisubpbl/450-Rewrite_Set_Inst.sml"
   218   ML_file "Minisubpbl/470-Check_elementwise-NEXT_STEP.sml"
   219   ML_file "Minisubpbl/490-nxt-Check_Postcond.sml"
   220   ML_file "Minisubpbl/500-met-sub-to-root.sml"
   221   ML_file "Minisubpbl/530-error-Check_Elementwise.sml"
   222   ML_file "Minisubpbl/600-postcond-NEXT_STEP.sml"
   223   ML_file "Minisubpbl/600-postcond.sml"
   224   ML_file "Minisubpbl/700-interSteps.sml"
   225   ML_file "Minisubpbl/710-interSteps-short.sml"
   226   ML_file "Minisubpbl/790-complete-NEXT_STEP.sml"
   227   ML_file "Minisubpbl/790-complete.sml"
   228   ML_file "Minisubpbl/800-append-on-Frm.sml"
   229 
   230 subsection \<open>further functionality alongside batch build sequence\<close>
   231   ML_file "MathEngBasic/thmC.sml"
   232   ML_file "MathEngBasic/rewrite.sml"
   233   ML_file "MathEngBasic/tactic.sml"
   234 (*ML_file "MathEngBasic/ctree.sml"  TOODOO Isa hangs ?!?*)
   235   ML_file "MathEngBasic/calculation.sml"
   236 
   237   ML_file "Specify/formalise.sml"
   238   ML_file "Specify/o-model.sml"
   239   ML_file "Specify/i-model.sml"
   240   ML_file "Specify/pre-conditions.sml"
   241   ML_file "Specify/p-model.sml"
   242   ML_file "Specify/m-match.sml"
   243   ML_file "Specify/refine.sml"        (* requires setup from refine.thy *)
   244   ML_file "Specify/test-out.sml"
   245   ML_file "Specify/specify-step.sml"
   246   ML_file "Specify/specification.sml"
   247   ML_file "Specify/cas-command.sml"
   248   ML_file "Specify/p-spec.sml"
   249   ML_file "Specify/specify.sml"
   250   ML_file "Specify/step-specify.sml"
   251 
   252   ML_file "Interpret/istate.sml"
   253   ML_file "Interpret/sub-problem.sml"
   254   ML_file "Interpret/error-pattern.sml"
   255   ML_file "Interpret/li-tool.sml"
   256   ML_file "Interpret/lucas-interpreter.sml"
   257 ML \<open>
   258 \<close> ML \<open>
   259 (* Title:  "Interpret/lucas-interpreter.sml"
   260    Author: Walther Neuper
   261    (c) due to copyright terms
   262 *)
   263 
   264 "-----------------------------------------------------------------------------------------------";
   265 "table of contents -----------------------------------------------------------------------------";
   266 "-----------------------------------------------------------------------------------------------";
   267 "----------- Take as 1st stac in program -------------------------------------------------------";
   268 "----------- re-build: fun locate_input_tactic -------------------------------------------------";
   269 "----------- fun locate_input_tactic Helpless, NOT applicable ----------------------------------";
   270 "----------- re-build: fun find_next_step, mini ------------------------------------------------";
   271 "----------- re-build: fun locate_input_term ---------------------------------------------------";
   272 "-----------------------------------------------------------------------------------------------";
   273 "-----------------------------------------------------------------------------------------------";
   274 "-----------------------------------------------------------------------------------------------";
   275 
   276 "----------- Take as 1st stac in program -------------------------------------------------------";
   277 "----------- Take as 1st stac in program -------------------------------------------------------";
   278 "----------- Take as 1st stac in program -------------------------------------------------------";
   279 "compare --- Apply_Method with initial Take by Step.do_next --- in test/../step-solve ----------";
   280 val p = e_pos'; val c = []; 
   281 val (p,_,f,nxt,_,pt) = 
   282     CalcTreeTEST 
   283         [(["functionTerm (x \<up> 2 + 1)", "integrateBy x", "antiDerivative FF"],
   284           ("Integrate", ["integrate", "function"], ["diff", "integration"]))];
   285 val (p,_,f,nxt,_,pt) = me nxt p c pt; (*nxt = ("Tac ", ...) --> Add_Given...*)
   286 val (p,_,f,nxt,_,pt) = me nxt p c pt;
   287 val (p,_,f,nxt,_,pt) = me nxt p c pt;
   288 val (p,_,f,nxt,_,pt) = me nxt p c pt;
   289 val (p,_,f,nxt,_,pt) = me nxt p c pt;
   290 val (p,_,f,nxt,_,pt) = me nxt p c pt;
   291 val (p,_,f,nxt,_,pt) = me nxt p c pt;
   292 case nxt of (Apply_Method ["diff", "integration"]) => ()
   293           | _ => error "integrate.sml -- me method [diff,integration] -- spec";
   294 "----- step 8: returns nxt = Rewrite_Set_Inst ([\"(''bdv'', x)\"],\"integration\")";
   295 
   296 "~~~~~ fun me, args:"; val (tac, (p:pos'), _, (pt:ctree)) = (nxt, p, c, pt);
   297 "~~~~~ fun Step.by_tactic, args:"; val (tac, (ptp as (pt, p))) = (tac, (pt,p));
   298 val Applicable.Yes m = Step.check tac (pt, p);
   299  (*if*) Tactic.for_specify' m; (*false*)
   300 "~~~~~ fun loc_solve_ , args:"; val (m, (pt,pos)) = (m, ptp);
   301 
   302 "~~~~~ fun Step_Solve.by_tactic , args:"; val (m as Apply_Method' (mI, _, _, ctxt), (pt, (pos as (p,_))))
   303   = (m, (pt, pos));
   304       val {srls, ...} = MethodC.from_store mI;
   305       val itms = case get_obj I pt p of
   306         PblObj {meth=itms, ...} => itms
   307       | _ => error "solve Apply_Method: uncovered case get_obj"
   308       val thy' = get_obj g_domID pt p;
   309       val thy = ThyC.get_theory thy';
   310       val srls = LItool.get_simplifier (pt, pos)
   311       val (is, env, ctxt, sc) = case LItool.init_pstate srls ctxt itms mI of
   312         (is as Istate.Pstate {env, ...}, ctxt, sc) =>  (is, env, ctxt, sc)
   313       | _ => error "solve Apply_Method: uncovered case init_pstate";
   314 (*+*)pstate2str (the_pstate is) = "([\"\n(f_f, x \<up> 2 + 1)\",\"\n(v_v, x)\"], [], empty, NONE, \n??.empty, ORundef, false, true)";
   315       val ini = LItool.implicit_take sc env;
   316       val p = lev_dn p;
   317 
   318       val NONE = (*case*) ini (*of*);
   319             val Next_Step (is', ctxt', m') =
   320               LI.find_next_step sc (pt, (p, Res)) is ctxt;
   321 (*+*)pstate2str (the_pstate is') = "([\"\n(f_f, x \<up> 2 + 1)\",\"\n(v_v, x)\"], [R,L,R], empty, NONE, \nIntegral x \<up> 2 + 1 D x, ORundef, false, false)";
   322   val Safe_Step (_, _, Take' _) = (*case*)
   323            locate_input_tactic sc (pt, (p, Res)) is' ctxt' m' (*of*);
   324 "~~~~~ fun locate_input_tactic , args:"; val ((Prog prog), cstate, istate, ctxt, tac)
   325   = (sc, (pt, (p, Res)), is', ctxt', m');
   326 
   327     (*case*) scan_to_tactic1 (prog, (cstate, ctxt, tac)) istate (*of*);
   328 "~~~~~ fun scan_to_tactic1 , args:"; val ((prog, (cctt as ((_, p), _, _))), (Istate.Pstate (ist as {path, ...})))
   329   = ((prog, (cstate, ctxt, tac)), istate);
   330     (*if*) path = [] orelse ((last_elem o fst) p = 0 andalso snd p = Res) (*then*);
   331 
   332   val Accept_Tac1 (_, _, Take' _) =
   333        scan_dn1 cctt (ist |> set_path [R] |> set_or ORundef) (Program.body_of prog);
   334 "~~~~~ fun scan_dn1 , args:"; val (xxx, ist, (Const ("HOL.Let", _) $ e $ (Abs (id, T, b))))
   335   = (cctt, (ist |> set_path [R] |> set_or ORundef), (Program.body_of prog));
   336 
   337 (*+*) if UnparseC.term e = "Take (Integral f_f D v_v)" then () else error "scan_dn1 Integral changed";
   338 
   339     (*case*)
   340            scan_dn1 xxx (ist |> path_down [L, R]) e (*of*);
   341     (*======= end of scanning tacticals, a leaf =======*)
   342 "~~~~~ fun scan_dn1 , args:"; val (((pt, p), ctxt, tac), (ist as {eval, or, ...}), t)
   343   = (xxx, (ist |> path_down [L, R]), e);
   344 val (Program.Tac stac, a') = check_leaf "locate" ctxt eval (get_subst ist) t;
   345 
   346 
   347 
   348 "----------- re-build: fun locate_input_tactic -------------------------------------------------";
   349 "----------- re-build: fun locate_input_tactic -------------------------------------------------";
   350 "----------- re-build: fun locate_input_tactic -------------------------------------------------";
   351 val fmz = ["equality (x+1=(2::real))", "solveFor x", "solutions L"];
   352 val (dI',pI',mI') = ("Test", ["sqroot-test", "univariate", "equation", "test"],
   353    ["Test", "squ-equ-test-subpbl1"]);
   354 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   355 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   356 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   357 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   358 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   359 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   360 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   361 (*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = (_, Apply_Method ["Test", "squ-equ-test-subpbl1"])*);
   362 (*[1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = ("Rewrite_Set", Rewrite_Set "norm_equation")*)
   363 
   364 (*//------------------ begin step into ------------------------------------------------------\\*)
   365 (*[1], Res*)val (p'''''_''',_,f,nxt'''''_''',_,pt'''''_''') = me nxt p [] pt; (*nxt = Rewrite_Set "Test_simplify"*)
   366 
   367 "~~~~~ fun me , args:"; val (tac, p, _(*NEW remove*), pt) = (nxt, p, [], pt);
   368 
   369     (** )val (***)xxxx(***) ( *("ok", (_, _, (pt, p))) =( **)  (*case*)
   370       Step.by_tactic tac (pt,p) (*of*);
   371 "~~~~~ fun by_tactic , args:"; val (tac, (ptp as (pt, p))) = (tac, (pt,p));
   372       val Applicable.Yes m = (*case*) Solve_Step.check tac (pt, p) (*of*);
   373       (*if*) Tactic.for_specify' m; (*false*)
   374 
   375     (** )val (***)xxxxx_x(***) ( *(msg, cs') =( **)
   376 Step_Solve.by_tactic m ptp;
   377 "~~~~~ fun by_tactic , args:"; val (m, (pt, po as (p, p_))) = (m, ptp);
   378 (*+*)val (pt'''''_', (p'''''_', p_'''''_')) = (pt, (p, p_));
   379     (*if*) MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p); (*else*)
   380 	      val thy' = get_obj g_domID pt (par_pblobj pt p);
   381 	      val (is, sc) = LItool.resume_prog thy' (p,p_) pt;
   382 
   383      locate_input_tactic sc (pt, po) (fst is) (snd is) m;
   384 "~~~~~ fun locate_input_tactic , args:"; val (Prog prog, cstate, istate, ctxt, tac)
   385     = (sc, (pt, po), (fst is), (snd is), m);
   386       val srls = get_simplifier cstate;
   387 
   388  (** )val Accept_Tac1 ((is as (_,_,_,_,_,strong_ass), ctxt, ss as((tac', _, ctree, pos', _) :: _))) =( **)
   389   (*case*) scan_to_tactic1 (prog, (cstate, ctxt, tac)) istate (*of*);
   390 "~~~~~ fun scan_to_tactic1 , args:"; val ((prog, (cctt as ((_, p), _, _))), (Istate.Pstate (ist as {path, ...})))
   391   = ((prog, (cstate, ctxt, tac)), istate);
   392     (*if*) path = [] orelse ((last_elem o fst) p = 0 andalso snd p = Res) (*then*);
   393 
   394     (** )val xxxxx_xx = ( **)
   395            scan_dn1 cctt (ist |> set_path [R] |> set_or ORundef) (Program.body_of prog);
   396 "~~~~~ fun scan_dn1 , args:"; val (xxx, ist, (Const ("HOL.Let", _) $ e $ (Abs (id, T, b))))
   397   = (cctt, (ist |> set_path [R] |> set_or ORundef), (Program.body_of prog));
   398 
   399   (*case*) scan_dn1 xxx (ist |> path_down [L, R]) e (*of*);
   400 "~~~~~ fun scan_dn1 , args:"; val ((xxx as (cstate, _, _)), ist, (Const ("Tactical.Chain"(*1*), _) $ e1 $ e2 $ a))
   401   = (xxx, (ist |> path_down [L, R]), e);
   402 
   403   (*case*) scan_dn1 xxx (ist |> path_down_form ([L, L, R], a)) e1 (*of*);
   404 "~~~~~ fun scan_dn1 , args:"; val (xxx, ist, (Const ("Tactical.Try"(*2*), _) $ e))
   405   = (xxx, (ist |> path_down_form ([L, L, R], a)), e1);
   406 
   407   (*case*) scan_dn1 xxx (ist |> path_down [R]) e (*of*);
   408     (*======= end of scanning tacticals, a leaf =======*)
   409 "~~~~~ fun scan_dn1 , args:"; val (((pt, p), ctxt, tac), (ist as {env, eval, or, ...}), t)
   410   = (xxx, (ist |> path_down [R]), e);
   411     val (Program.Tac stac, a') =
   412       (*case*) check_leaf "locate" ctxt eval (get_subst ist) t (*of*);
   413     val LItool.Associated (m, v', ctxt) =
   414       (*case*) associate pt ctxt (m, stac) (*of*);
   415 
   416        Accept_Tac1 (ist |> set_subst_true  (a', v'), ctxt, m)  (*return value*);
   417 "~~~~~ from scan_dn1 to scan_to_tactic1 return val:"; val (xxxxx_xx)
   418   = (Accept_Tac1 (ist |> set_subst_true  (a', v'), ctxt, m));
   419 
   420 "~~~~~ from scan_to_tactic1 to fun locate_input_tactic return val:"; val Accept_Tac1 ((ist as {assoc, ...}), ctxt, tac')
   421   = (Accept_Tac1 (ist |> set_subst_true  (a', v'), ctxt, m));
   422         (*if*) LibraryC.assoc (*then*);
   423 
   424        Safe_Step (Istate.Pstate ist, ctxt, tac')  (*return value*);
   425 "~~~~~ from locate_input_tactic to fun Step_Solve.by_tactic return:"; val Safe_Step (istate, ctxt, tac)
   426   = (*xxxxx_xx*)(**)Safe_Step (Istate.Pstate ist, ctxt, tac')(**);
   427 
   428 (*+*)val (pt, po as (p, p_)) = (pt'''''_', (p'''''_', p_'''''_')); (* from begin of by_tactic *)
   429                   val (p'', _, _,pt') =
   430                     Step.add tac (istate, ctxt) (pt, (lev_on p, Pbl));
   431             (*in*)
   432 
   433          	  	    ("ok", ([(Tactic.input_from_T tac, tac, (p'', (istate, ctxt)))],
   434                     [(*ctree NOT cut*)], (pt', p'')))  (*return value*);
   435 "~~~~~ from Step_Solve.by_tactic \<longrightarrow> Step.by_tactic return:"; val ((msg, cs' : Calc.state_post))
   436   =               ("ok", ([(Tactic.input_from_T tac, tac, (p'', (istate, ctxt)) )],
   437                     [(*ctree NOT cut*)], (pt', p'')));
   438 
   439 "~~~~~ from Step.by_tactic to me return:"; val (("ok", (_, _, (pt, p)))) = (*** )xxxx( ***) ("ok", cs');
   440 	  val (_, ts) =
   441 	    (case Step.do_next p ((pt, Pos.e_pos'), []) of
   442 		    ("ok", (ts as (_, _, _) :: _, _, _)) => ("", ts)
   443 	    | ("helpless", _) => ("helpless: cannot propose tac", [])
   444 	    | ("no-fmz-spec", _) => error "no-fmz-spec"
   445 	    | ("end-of-calculation", (ts, _, _)) => ("", ts)
   446 	    | _ => error "me: uncovered case")
   447 	      handle ERROR msg => raise ERROR msg
   448 	  val tac = 
   449       case ts of 
   450         tacis as (_::_) => let val (tac, _, _) = last_elem tacis in tac end 
   451 		  | _ => if p = ([], Pos.Res) then Tactic.End_Proof' else Tactic.Empty_Tac;
   452 
   453    (p, [] : NEW, TESTg_form (pt, p), (Tactic.tac2IDstr tac, tac), Celem.Sundef, pt);
   454 "~~~~~ from me to TOOPLEVEL return:"; val (p,_,f,nxt,_,pt)
   455    = (*** )xxx( ***) (p, [] : NEW, TESTg_form (pt, p), (Tactic.tac2IDstr tac, tac), Celem.Sundef, pt);
   456 
   457 (*//--------------------- check results from modified me ----------------------------------\\*)
   458 if p = ([2], Res) andalso
   459   pr_ctree pr_short pt = ".    ----- pblobj -----\n1.   x + 1 = 2\n2.   x + 1 = 2\n"
   460 then
   461   (case nxt of ("Rewrite_Set", Rewrite_Set "Test_simplify") => ()
   462    | _ => error "")
   463 else error "check results from modified me CHANGED";
   464 (*\\--------------------- check results from modified me ----------------------------------//*)
   465 
   466 "~~~~~ from me to TOPLEVEL return:"; val (p,_,f,nxt,_,pt) = (*** )xxx( ***) (**)(p, 000, f, nxt, 000, pt)(**);
   467 (*\\------------------ end step into --------------------------------------------------------//*)
   468 
   469 (*[3], Res*)val (p,_,f,nxt,_,pt) = me nxt'''''_''' p'''''_''' [] pt'''''_'''; (*nxt = Subproblem ("Test", ["LINEAR", "univariate", "equation", "test"])*)
   470 (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Model_Problem*)
   471 (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Add_Given "equality (- 1 + x = 0)"*)
   472 (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Add_Given "solveFor x"*)
   473 (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Add_Find "solutions x_i"*)
   474 (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Specify_Theory "Test"*)
   475 (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Specify_Problem ["LINEAR", "univariate", "equation", "test"]*)
   476 (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Specify_Method ["Test", "solve_linear"]*)
   477 (*[3], Met*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Apply_Method ["Test", "solve_linear"]*)
   478 (*[3, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Rewrite_Set_Inst (["(''bdv'', x)"], "isolate_bdv")*)
   479 (*[3, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Rewrite_Set "Test_simplify"*)
   480 (*[3, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Check_Postcond ["LINEAR", "univariate", "equation", "test"]*)
   481 (*[3], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Check_elementwise "Assumptions"*)
   482 (*[4], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Check_Postcond ["sqroot-test", "univariate", "equation", "test"]*)
   483 (*[], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = End_Proof'*)
   484 
   485 (*/--------------------- final test ----------------------------------\\*)
   486 if p = ([], Res) andalso f2str f = "[x = 1]" andalso pr_ctree pr_short pt =
   487   ".    ----- pblobj -----\n" ^
   488   "1.   x + 1 = 2\n" ^
   489   "2.   x + 1 + - 1 * 2 = 0\n" ^
   490   "3.    ----- pblobj -----\n" ^
   491   "3.1.   - 1 + x = 0\n" ^
   492   "3.2.   x = 0 + - 1 * - 1\n" ^
   493   "4.   [x = 1]\n"
   494 then case nxt of End_Proof' => () | _ => error "re-build: fun locate_input_tactic changed 1"
   495 else error "re-build: fun locate_input_tactic changed 2";
   496 
   497 
   498 \<close> ML \<open>
   499 "----------- fun locate_input_tactic Helpless, NOT applicable ----------------------------------";
   500 "----------- fun locate_input_tactic Helpless, NOT applicable ----------------------------------";
   501 "----------- fun locate_input_tactic Helpless, NOT applicable ----------------------------------";
   502 (*cp from -- try fun applyTactics ------- *)
   503 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(["Term (5*e + 6*f - 8*g - 9 - 7*e - 4*f + 10*g + 12)",
   504 	    "normalform N"],
   505 	   ("PolyMinus",["plus_minus", "polynom", "vereinfachen"],
   506 	    ["simplification", "for_polynomials", "with_minus"]))];
   507 val (p,_,f,nxt,_,pt) = me nxt p [] pt; val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   508 val (p,_,f,nxt,_,pt) = me nxt p [] pt; val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   509 val (p,_,f,nxt,_,pt) = me nxt p [] pt; val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   510 \<close> ML \<open>
   511 f
   512 \<close> ML \<open>
   513 (*[1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Rewrite_Set "ordne_alphabetisch"*)
   514 \<close> ML \<open>
   515 val Test_Out.FormKF "5 * e + 6 * f - 8 * g - 9 - 7 * e - 4 * f + 10 * g + 12" = f
   516 \<close> ML \<open>
   517 (*[1], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Rewrite_Set "fasse_zusammen"*)
   518 
   519 \<close> ML \<open>
   520 f
   521 \<close> ML \<open>
   522 map Tactic.input_to_string (specific_from_prog pt p)
   523 \<close> ML \<open>
   524 (*+*)if map Tactic.input_to_string (specific_from_prog pt p) =
   525    ["Rewrite (\"subtrahiere_x_plus_minus\", \"\<lbrakk>?l is_const; ?m is_const\<rbrakk>\n\<Longrightarrow> ?x + ?m * ?v - ?l * ?v = ?x + (?m - ?l) * ?v\")",
   526     "Rewrite (\"subtrahiere_x_minus_plus\", \"\<lbrakk>?l is_const; ?m is_const\<rbrakk>\n\<Longrightarrow> ?x - ?m * ?v + ?l * ?v = ?x + (- ?m + ?l) * ?v\")"]
   527   then () else error "specific_from_prog ([1], Res) CHANGED";
   528 (*[2], Res*)val ("ok", (_, _, ptp as (pt, p))) = Step.by_tactic (hd (specific_from_prog pt p)) (pt, p);
   529 
   530 \<close> ML \<open>
   531 (*+*)if map Tactic.input_to_string (specific_from_prog pt p) =
   532    ["Rewrite (\"tausche_minus\", \"\<lbrakk>?b ist_monom; ?a kleiner ?b\<rbrakk>\n\<Longrightarrow> ?b - ?a = - ?a + ?b\")", "Rewrite (\"tausche_plus_minus\", \"?b kleiner ?c \<Longrightarrow> ?a + ?c - ?b = ?a - ?b + ?c\")",
   533     "Rewrite (\"subtrahiere_x_plus_minus\", \"\<lbrakk>?l is_const; ?m is_const\<rbrakk>\n\<Longrightarrow> ?x + ?m * ?v - ?l * ?v = ?x + (?m - ?l) * ?v\")",
   534     "Rewrite (\"subtrahiere_x_minus_plus\", \"\<lbrakk>?l is_const; ?m is_const\<rbrakk>\n\<Longrightarrow> ?x - ?m * ?v + ?l * ?v = ?x + (- ?m + ?l) * ?v\")", "Calculate MINUS"]
   535   then () else error "specific_from_prog ([1], Res) CHANGED";
   536 (* = ([3], Res)*)val ("ok", (_, _, ptp as (pt, p))) = Step.by_tactic (hd (specific_from_prog pt p)) (pt, p);
   537 
   538 \<close> ML \<open>
   539 (*//----------------- exception PTREE "get_obj f EmptyPtree" raised --------------------------\\*)
   540 (** )val ("ok", (_, _, ptp as (pt, p))) =( **)
   541       Step.by_tactic (hd (specific_from_prog pt p)) (pt, p);
   542 "~~~~~ fun by_tactic , args:"; val (tac, (ptp as (pt, p))) = (hd (specific_from_prog pt p), (pt, p));
   543       val Applicable.Yes m = (*case*) Solve_Step.check tac (pt, p) (*of*);
   544       (*if*) Tactic.for_specify' m; (*false*)
   545 
   546 Step_Solve.by_tactic m (pt, p);
   547 "~~~~~ fun by_tactic , args:"; val (m, (pt, po as (p, p_))) = (m, (pt, p));
   548     (*if*) MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p) (*else*);
   549 	      val thy' = get_obj g_domID pt (par_pblobj pt p);
   550 	      val (is, sc) = LItool.resume_prog thy' (p,p_) pt;
   551 
   552   (*case*) locate_input_tactic sc (pt, po) (fst is) (snd is) m (*of*);
   553 "~~~~~ fun locate_input_tactic , args:"; val ((Rule.Prog prog), (cstate as (pt, (*?*)pos(*?*))), istate, ctxt, tac)
   554   = (sc, (pt, po), (fst is), (snd is), m);
   555       val srls = LItool.get_simplifier cstate (*TODO: shift into Istate.T*);
   556 
   557   (*case*) scan_to_tactic1 (prog, (cstate, ctxt, tac)) istate (*of*);
   558 "~~~~~ fun scan_to_tactic1 , args:"; val ((prog, (cctt as ((_, p), _, _))), (Istate.Pstate (ist as {path, ...})))
   559   = ((prog, (cstate, ctxt, tac)), istate);
   560     (*if*) path = [] orelse ((last_elem o fst) p = 0 andalso snd p = Res) (*else*);
   561 
   562            go_scan_up1 (prog, cctt) ist;
   563 "~~~~~ fun go_scan_up1 , args:"; val ((pcct as (prog, _)), (ist as {path, ...}))
   564   = ((prog, cctt), ist);
   565   (*if*) 1 < length path (*then*);
   566 
   567            scan_up1 pcct (ist |> path_up) (TermC.sub_at (path_up' path) prog);
   568 "~~~~~ and scan_up1 , args:"; val (pcct, ist, (Const ("Tactical.Try"(*2*), _) $ _))
   569   = (pcct, (ist |> path_up), (TermC.sub_at (path_up' path) prog));
   570 
   571            go_scan_up1 pcct ist;
   572 "~~~~~ and go_scan_up1 , args:"; val ((pcct as (prog, _)), (ist as {path, ...}))
   573   = (pcct, ist);
   574   (*if*) 1 < length path (*then*);
   575 
   576            scan_up1 pcct (ist |> path_up) (TermC.sub_at (path_up' path) prog);
   577 "~~~~~ and scan_up1 , args:"; val ((pcct as (prog, cct as (cstate, _, _))), ist,
   578     (Const ("Tactical.Chain"(*3*), _) $ _ ))
   579   = (pcct, (ist |> path_up), (TermC.sub_at (path_up' path) prog));
   580        val e2 = check_Seq_up ist prog
   581 ;
   582   (*case*) scan_dn1 cct (ist |> path_up_down [R] |> set_or ORundef) e2 (*of*);
   583 "~~~~~ fun scan_dn1 , args:"; val (yyy, ist, (Const ("Tactical.Chain"(*2*), _) $ e1 $ e2))
   584   = (cct, (ist |> path_up_down [R] |> set_or ORundef), e2);
   585 
   586   (*case*) scan_dn1 cct (ist |> path_down [L, R]) e1 (*of*);
   587 "~~~~~ fun scan_dn1 , args:"; val (xxx, ist, (Const ("Tactical.Try"(*2*), _) $ e))
   588   = (cct, (ist |> path_down [L, R]), e1);
   589 
   590   (*case*) scan_dn1 cct (ist |> path_down [R]) e (*of*);
   591     (*======= end of scanning tacticals, a leaf =======*)
   592 "~~~~~ fun scan_dn1 , args:"; val ((cct as (_, ctxt, _)), (ist as {eval, ...}), t)
   593   = (cct, (ist |> path_down [R]), e);
   594     (*if*) Tactical.contained_in t (*else*);
   595   val (Program.Tac prog_tac, form_arg) = (*case*)
   596     LItool.check_leaf "locate" ctxt eval (get_subst ist) t (*of*);
   597 
   598            check_tac1 cct ist (prog_tac, form_arg);
   599 "~~~~~ fun check_tac1 , args:"; val (((pt, p), ctxt, tac), (ist as {act_arg, or, ...}), (prog_tac, form_arg)) =
   600   (cct, ist, (prog_tac, form_arg));
   601 val LItool.Not_Associated = (*case*)
   602   LItool.associate pt ctxt (tac, prog_tac) (*of*);
   603      val _(*ORundef*) = (*case*) or (*of*);
   604      val Applicable.Yes m' =
   605           (*case*) Solve_Step.check (LItool.tac_from_prog pt (ThyC.get_theory "Isac_Knowledge") prog_tac) (pt, p) (*of*);
   606 
   607   Reject_Tac1 (ist |> set_subst_false (form_arg, Tactic.result m'), ctxt, tac)
   608           (*return from check_tac1*);
   609 "~~~~~ from fun check_tac1 \<longrightarrow>fun scan_dn1 \<longrightarrow>fun scan_dn1  \<longrightarrow>fun locate_input_tactic , return:"; val (Reject_Tac1 _) =
   610   (Reject_Tac1 (ist |> set_subst_false (form_arg, Tactic.result m'), ctxt, tac));
   611 
   612 (*/----- original before child of 7e314dd233fd -------------------------------------------------\* )
   613     val (Program.Tac prog_tac, form_arg) = (*case*) check_leaf "locate" ctxt eval (get_subst ist) t (*of*);
   614       val Not_Associated = (*case*) associate pt ctxt (tac, stac) (*of*);
   615       val ORundef = (*case*) or (*of*);
   616   val Applicable.No "norm_equation not applicable" =    
   617       (*case*) Solve_Step.check (LItool.tac_from_prog pt (ThyC.get_theory "Isac_Knowledge") stac) (pt, p) (*of*);
   618 
   619    (Term_Val1 act_arg) (* return value *);
   620 
   621 val Rewrite' ("PolyMinus", "tless_true", _, _, ("tausche_minus",_ (*"?b ist_monom \<Longrightarrow> ?a kleiner ?b \<Longrightarrow> ?b - ?a = - ?a + ?b"*)),
   622    t, (res, asm)) = m;
   623 
   624 if pstate2str ist =
   625   "([\"\n(t_t, 5 * e + 6 * f - 8 * g - 9 - 7 * e - 4 * f + 10 * g + 12)\"], [R,L,R,R,L,R,R], empty, SOME t_t, \n" ^
   626   "- (8 * g) + (- 9 + 12 + 5 * e - 7 * e + (6 - 4) * f) + 10 * g, ORundef, true, false)"
   627 andalso
   628   UnparseC.term t = "- (8 * g) + (- 9 + 12 + 5 * e - 7 * e + (6 - 4) * f) + 10 * g"
   629 andalso
   630   UnparseC.term res = "- (8 * g) + (- 9 + 12 + 5 * e - 7 * e + (- 4 + 6) * f) + 10 * g"
   631 andalso
   632   UnparseC.terms asm = "[\"4 kleiner 6\",\"6 ist_monom\"]"
   633 then () else error "locate_input_tactic Helpless, but applicable CHANGED";
   634 ( *\----- original before child of 7e314dd233fd -------------------------------------------------/*)
   635 
   636 
   637 \<close> ML \<open>
   638 "----------- re-build: fun find_next_step, mini ------------------------------------------------";
   639 "----------- re-build: fun find_next_step, mini ------------------------------------------------";
   640 "----------- re-build: fun find_next_step, mini ------------------------------------------------";
   641 val fmz = ["Term (a + a ::real)", "normalform n_n"];
   642 val (dI',pI',mI') = ("Poly",["polynomial", "simplification"],["simplification", "for_polynomials"]);
   643 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   644 (*[], Pbl*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, Pos.e_pos'), []);(*Model_Problem*)
   645 (*[], Pbl*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, Pos.e_pos'), []);(*Specify_Theory "Poly"*)
   646 (*[], Pbl*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, Pos.e_pos'), []);(*Specify_Problem ["polynomial", "simplification"]*)
   647 (*[], Met*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, e_pos'), []);(*Specify_Method  ["simplification", "for_polynomials"]*)
   648 (*1], Frm*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, e_pos'), []);(*Apply_Method ["simplification", "for_polynomials"]*)
   649 (*[1], Res*)val (_, ([(tac'''''_', _, _)], _, (pt'''''_', p'''''_'))) =
   650 
   651       Step.do_next p ((pt, e_pos'), []);(*Rewrite_Set "norm_Poly"*)
   652 \<close> ML \<open>
   653 (*//------------------ go into 1 ------------------------------------------------------------\\*)
   654 "~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis))
   655   = (p, ((pt, e_pos'), []));
   656   val pIopt = Ctree.get_pblID (pt, ip);
   657     (*if*)  ip = ([], Res) (*else*);
   658       val _ = (*case*) tacis (*of*);
   659       val SOME _ = (*case*) pIopt (*of*);
   660       (*if*) member op = [Pos.Pbl, Pos.Met] p_ (*else*);
   661 
   662 val ("ok", ([(Rewrite_Set "norm_Poly", _, _)], _, (_, ([1], Res)))) =
   663 Step_Solve.do_next (pt, ip);
   664 "~~~~~ and do_next , args:"; val (ptp as (pt, pos as (p, p_))) =  (pt, ip);
   665     (*if*) MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p) (*else*);
   666         val thy' = get_obj g_domID pt (par_pblobj pt p);
   667 	      val ((ist, ctxt), sc) = LItool.resume_prog thy' (p,p_) pt;
   668 
   669 val Next_Step (_, _, Rewrite_Set' ("Poly", _, Rule_Set.Sequence {id = "norm_Poly", ...}, _, _)) =
   670            LI.find_next_step sc (pt, pos) ist ctxt (*of*);
   671 "~~~~~ fun find_next_step , args:"; val ((Rule.Prog prog), (ptp as(pt, (p, _))), (Pstate ist), ctxt)
   672   = (sc, (pt, pos), ist, ctxt);
   673 
   674 val Accept_Tac (Rewrite_Set' ("Poly", _, Rule_Set.Sequence {id = "norm_Poly", ...}, _, _), _, _) =
   675   (*case*) scan_to_tactic (prog, (ptp, ctxt)) (Pstate ist) (*of*);
   676 "~~~~~ fun scan_to_tactic , args:"; val ((prog, cc), (Pstate (ist as {path, ...})))
   677   = ((prog, (ptp, ctxt)), (Pstate ist));
   678   (*if*) path = [] (*then*);
   679 
   680 val Accept_Tac (Rewrite_Set' ("Poly", _, Rule_Set.Sequence {id = "norm_Poly", ...}, _, _), _, _) =
   681             scan_dn cc (trans_scan_dn ist) (Program.body_of prog);
   682 "~~~~~ fun scan_dn , args:"; val ((cc as (_, ctxt)), (ist as {eval, ...}), t)
   683   = (cc, (trans_scan_dn ist), (Program.body_of prog));
   684     (*if*) Tactical.contained_in t (*else*);
   685       val (Program.Tac prog_tac, form_arg) = (*case*) LItool.check_leaf "next  " ctxt eval (get_subst ist) t (*of*);
   686 
   687 val Accept_Tac (Rewrite_Set' ("Poly", _, Rule_Set.Sequence {id = "norm_Poly", ...}, _, _), _, _) =
   688           check_tac cc ist (prog_tac, form_arg)  (*return from xxx*);
   689 "~~~~~ from fun scan_dn\<longrightarrow>fun scan_to_tactic\<longrightarrow>fun find_next_step, return:"; val (Accept_Tac (tac, ist, ctxt))
   690   = (check_tac cc ist (prog_tac, form_arg));
   691 
   692     Next_Step (Pstate ist, Tactic.insert_assumptions tac ctxt, tac)  (*return from find_next_step*);
   693 "~~~~~ from fun find_next_step\<longrightarrow>and do_next\<longrightarrow>fun zzz, return:"; val (Next_Step (ist, ctxt, tac))
   694   = (Next_Step (Pstate ist, Tactic.insert_assumptions tac ctxt, tac));
   695 
   696            LI.by_tactic tac (ist, Tactic.insert_assumptions tac ctxt) ptp  (*return from and do_next*);
   697 "~~~~~ from and do_next\<longrightarrow>fun do_next\<longrightarrow>toplevel, return:"; val  (_, ([(tac''''', _, _)], _, (pt''''', p''''')))
   698   = (LI.by_tactic tac (ist, Tactic.insert_assumptions tac ctxt) ptp);
   699 (*\\------------------ end of go into 1 -----------------------------------------------------//*)
   700 
   701 (*[], Res*)val (_, ([(tac''''', _, _)], _, (pt''''', p'''''))) =
   702 
   703       Step.do_next p'''''_' ((pt'''''_', Pos.e_pos'), []);(* Check_Postcond ["polynomial", "simplification"]*)
   704 (*//------------------ go into 2 ------------------------------------------------------------\\*)
   705 "~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis))
   706   = (p''''', ((pt''''', e_pos'), []));
   707   val pIopt = Ctree.get_pblID (pt, ip);
   708     (*if*)  ip = ([], Res) (*else*);
   709       val _ = (*case*) tacis (*of*);
   710       val SOME _ = (*case*) pIopt (*of*);
   711       (*if*) member op = [Pos.Pbl, Pos.Met] p_ (*else*);
   712 
   713 val ("ok", ([(Check_Postcond ["polynomial", "simplification"], _, _)], _, (_, ([], Res)))) =
   714 Step_Solve.do_next (pt, ip);
   715 "~~~~~ and do_next , args:"; val (ptp as (pt, pos as (p, p_))) =  (pt, ip);
   716     (*if*) MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p) (*else*);
   717         val thy' = get_obj g_domID pt (par_pblobj pt p);
   718 	      val ((ist, ctxt), sc) = LItool.resume_prog thy' (p,p_) pt;
   719 
   720   (** )val End_Program (ist, tac) = 
   721  ( *case*) LI.find_next_step sc (pt, pos) ist ctxt (*of*);
   722 "~~~~~ fun find_next_step , args:"; val ((Rule.Prog prog), (ptp as(pt, (p, _))), (Pstate ist), ctxt)
   723   = (sc, (pt, pos), ist, ctxt);
   724 
   725 (*  val Term_Val (Const ("Groups.times_class.times", _) $ Free ("2", _) $ Free ("a", _))*)
   726   (** )val Term_Val prog_result =
   727  ( *case*) scan_to_tactic (prog, (ptp, ctxt)) (Pstate ist) (*of*);
   728 "~~~~~ fun scan_to_tactic , args:"; val ((prog, cc), (Pstate (ist as {path, ...})))
   729   = ((prog, (ptp, ctxt)), (Pstate ist));
   730   (*if*) path = [] (*else*);
   731 
   732            go_scan_up (prog, cc) (trans_scan_up ist |> set_found);
   733 "~~~~~ fun go_scan_up , args:"; val ((pcc as (sc, _)), (ist as {path, act_arg, found_accept, ...}))
   734   = ((prog, cc), (trans_scan_up ist(*|> set_found !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!! *)));
   735     (*if*) path = [R] (*then*);
   736       (*if*) found_accept = true (*then*);
   737 
   738       Term_Val act_arg (*return from go_scan_up*);
   739 "~~~~~ from fun go_scan_up\<longrightarrow>fun scan_to_tactic, return:"; val (Term_Val prog_result) = (Term_Val act_arg);
   740 
   741     Term_Val prog_result  (*return from scan_to_tactic*);
   742 "~~~~~ from fun scan_to_tactic\<longrightarrow>fun find_next_step, return:"; val (Term_Val prog_result) = (Term_Val prog_result);
   743     val (true, p', _) = (*case*) parent_node pt p (*of*);
   744               val (_, pblID, _) = get_obj g_spec pt p';
   745 
   746      End_Program (Pstate ist, Tactic.Check_Postcond' (pblID, prog_result))
   747      (*return from find_next_step*);
   748 "~~~~~ from fun find_next_step\<longrightarrow>and do_next\<longrightarrow>fun zzz, return:"; val (End_Program (ist, tac))
   749   = (End_Program (Pstate ist, Tactic.Check_Postcond' (pblID,prog_result)));
   750       val _ = (*case*) tac (*of*);
   751 
   752 val ("ok", ([(Check_Postcond ["polynomial", "simplification"], _, _)], _, (_, ([], Res))))
   753    = LI.by_tactic tac (ist, ctxt) ptp (*return from and do_next*);
   754 "~~~~~ from and do_next\<longrightarrow>top level, return:"; val (_, ([(tac''''', _, _)], _, (pt''''', p''''')))
   755   = (LI.by_tactic tac (ist, ctxt) ptp);
   756 (*\\------------------ end of go into 2 -----------------------------------------------------//*)
   757 
   758 (*[], Und*)val (msg, ([], _, (pt, p))) = Step.do_next p''''' ((pt''''', Pos.e_pos'), []);(**)
   759 
   760 Test_Tool.show_pt_tac pt; (*[
   761 ([], Frm), Simplify (a + a)
   762 . . . . . . . . . . Apply_Method ["simplification", "for_polynomials"],
   763 ([1], Frm), a + a
   764 . . . . . . . . . . Rewrite_Set "norm_Poly",
   765 ([1], Res), 2 * a
   766 . . . . . . . . . . Check_Postcond ["polynomial", "simplification"],
   767 ([], Res), 2 * a]*)
   768 
   769 (*/--- final test ---------------------------------------------------------------------------\\*)
   770 val (res, asm) = (get_obj g_result pt (fst p));
   771 if UnparseC.term res = "2 * a" andalso map UnparseC.term asm = []
   772 andalso p = ([], Und) andalso msg = "end-of-calculation"
   773 andalso pr_ctree pr_short pt = ".    ----- pblobj -----\n1.   a + a\n"
   774 then 
   775   case tac''''' of Check_Postcond ["polynomial", "simplification"] => () 
   776   | _ => error "re-build: fun find_next_step, mini 1"
   777 else error "re-build: fun find_next_step, mini 2"
   778 
   779 
   780 \<close> ML \<open>
   781 "----------- re-build: fun locate_input_term ---------------------------------------------------";
   782 "----------- re-build: fun locate_input_term ---------------------------------------------------";
   783 "----------- re-build: fun locate_input_term ---------------------------------------------------";
   784 (*cp from inform.sml
   785  ----------- appendFormula: on Res + late deriv ------------------------------------------------*)
   786 val fmz = ["equality (x+1=(2::real))", "solveFor x", "solutions L"];
   787 val (dI',pI',mI') = ("Test", ["sqroot-test", "univariate", "equation", "test"],
   788    ["Test", "squ-equ-test-subpbl1"]);
   789 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   790 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   791 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   792 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   793 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   794 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   795 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   796 (*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Apply_Method ["Test", "squ-equ-test-subpbl1"]*);
   797 
   798 (*[1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p [] pt;(*Rewrite_Set "norm_equation"*)
   799 (*+*)if f2str f = "x + 1 = 2" then () else error "locate_input_term at ([1], Frm) CHANGED";
   800 
   801 (*[1], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt;(*Rewrite_Set "Test_simplify"*)
   802 (*+*)if f2str f = "x + 1 + - 1 * 2 = 0" then () else error "locate_input_term at ([1], Frm) CHANGED";
   803 
   804 Test_Tool.show_pt_tac pt; (*[
   805 ([], Frm), solve (x + 1 = 2, x)
   806 . . . . . . . . . . Apply_Method ["Test", "squ-equ-test-subpbl1"],
   807 ([1], Frm), x + 1 = 2
   808 . . . . . . . . . . Rewrite_Set "norm_equation",
   809 ([1], Res), x + 1 + - 1 * 2 = 0             ///Check_Postcond..ERROR*)
   810 
   811 (*//---------- appendFormula 1 "x = 1" \<longrightarrow> Step_Solve.inform \<longrightarrow> LI.locate_input_term ----------\\*)
   812 "~~~~~ fun appendFormula , args:"; val ((*cI, *) ifo: TermC.as_string) = ((**) "x = 1");
   813     val cs = (*get_calc cI*) ((pt, p), [(*nxt, nxt_, (pos, (ist, ctxt))*)])
   814     val pos = (*get_pos cI 1*) p
   815 
   816 (*+*)val ptp''''' = (pt, p);
   817 (*+*)if snd ptp''''' = ([1], Res) then () else error "old_cs changed";
   818 (*+*)Test_Tool.show_pt_tac pt; (*[
   819 (*+*)([], Frm), solve (x + 1 = 2, x)
   820 (*+*). . . . . . . . . . Apply_Method ["Test", "squ-equ-test-subpbl1"],
   821 (*+*)([1], Frm), x + 1 = 2
   822 (*+*). . . . . . . . . . Rewrite_Set "norm_equation",
   823 (*+*)([1], Res), x + 1 + - 1 * 2 = 0      ///Check_Postcond*)
   824 
   825   val ("ok", cs' as (_, _, ptp')) =
   826     (*case*) Step.do_next pos cs (*of*);
   827 
   828 val ("ok", (_(*use in DG !!!*), [], ptp''''' as (pt''''', p'''''))) = (*case*)
   829      Step_Solve.by_term ptp' (encode ifo) (*of*);
   830 "~~~~~ fun Step_Solve.by_term , args:"; val ((pt, pos as (p, _)), istr)
   831   = (ptp', (encode ifo));
   832   val SOME f_in =
   833     (*case*) TermC.parse (ThyC.get_theory "Isac_Knowledge") istr (*of*);
   834   	  val f_in = Thm.term_of f_in
   835       val pos_pred = lev_back(*'*) pos
   836   	  val f_pred = Ctree.get_curr_formula (pt, pos_pred);
   837   	  val f_succ = Ctree.get_curr_formula (pt, pos);
   838       (*if*) f_succ = f_in (*else*);
   839   val NONE =
   840         (*case*) CAS_Cmd.input f_in (*of*);
   841 
   842 (*old* )     val (_, _, metID) = get_obj g_spec pt (par_pblobj pt p)
   843 (*old*)     val {scr = prog, ...} = MethodC.from_store metID
   844 (*old*)     val istate = get_istate_LI pt pos
   845 (*old*)     val ctxt = get_ctxt pt pos
   846   val LI.Found_Step (cstate'''''_', _(*istate*), _(*ctxt*)) = (*case*)
   847         LI.locate_input_term prog (pt, pos) istate ctxt f_in (*of*);
   848 "~~~~~ fun locate_input_term , args:"; val ((Rule.Prog _),  ((pt, pos) : Calc.T), (_ : Istate.T), (_ : Proof.context), tm)
   849   = (prog, (pt, pos), istate, ctxt, f_in);
   850 ( *old*)
   851 
   852 (*NEW*) LI.locate_input_term (pt, pos) f_in (*of*);
   853 "~~~~~ fun locate_input_term , args:"; val ((pt, pos), tm) = ((pt, pos), f_in);
   854    		val pos_pred = Pos.lev_back' pos (*f_pred ---"step pos cs"---> f_succ in appendFormula*)
   855 
   856   val ("ok", (_, _, cstate as (pt', pos'))) =
   857    		(*case*) compare_step ([], [], (pt, pos_pred)) tm (*of*);
   858 
   859 (*old* )
   860     Found_Step (cstate, get_istate_LI pt' pos', get_ctxt pt' pos')  (*return from locate_input_term*);
   861 ( *old*)
   862 (*NEW*)     Found_Step cstate (*return from locate_input_term*);
   863        (*LI.Found_Step ( *)cstate(*, _(*istate*), _(*ctxt*))( *return from locate_input_term*);
   864 "~~~~~ from fun locate_input_term\<longrightarrow>fun Step_Solve.by_term, return:"; val ("ok", (_(*use in DG !!!*), c, ptp as (_, p)))
   865   = (("ok" , ([], [], cstate (* already contains istate, ctxt *))));
   866 
   867     ("ok", ((*_ use in DG !!!,*) c, ptp(* as (_*), p))(*)*)(*return from Step_Solve.by_term*);
   868 "~~~~~ from fun Step_Solve.by_term\<longrightarrow>(fun appendFormula)!toplevel, return:"; val ("ok", (_(*use in DG !!!*), [], ptp''''' as (pt''''', p''''')))
   869   = ("ok", ([], [], ptp));
   870 
   871 (*fun me requires nxt...*)
   872     Step.do_next p''''' (ptp''''', []);
   873   val ("ok", ([(nxt'''''_' as Check_Postcond ["LINEAR", "univariate", "equation", "test"], _, _)], _,
   874     (pt'''''_', p'''''_'))) = Step.do_next p''''' (ptp''''', [])
   875 (*\\---------- appendFormula 1 "x = 1" \<longrightarrow> Step_Solve.inform \<longrightarrow> LI.locate_input_term ----------//*)
   876 
   877 (*//----- REPLACED BY appendFormula 1 "x = 1" \<longrightarrow> Step_Solve.inform \<longrightarrow> LI.locate_input_term -----\\* )
   878  (*[2], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Subproblem ("Test", ["LINEAR", "univariate", "equation", "test"])*)
   879  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Model_Problem*)
   880  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Add_Given "equality (- 1 + x = 0)"*)
   881  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Add_Given "solveFor x"*)
   882  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Add_Find "solutions x_i"*)
   883  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Specify_Theory "Test"*)
   884  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Specify_Problem ["LINEAR", "univariate", "equation", "test"]*)
   885  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Specify_Method ["Test", "solve_linear"]*)
   886  (*[3], Met*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Apply_Method ["Test", "solve_linear"]*)
   887  (*[3, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Rewrite_Set_Inst (["(''bdv'', x)"], "isolate_bdv")*)
   888  (*[3, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Rewrite_Set "Test_simplify"*)
   889  (*[3, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Check_Postcond ["LINEAR", "univariate", "equation", "test"]*)
   890 ( *\\----- REPLACED BY appendFormula 1 "x = 1" \<longrightarrow> Step_Solve.inform \<longrightarrow> LI.locate_input_term -----//*)
   891 
   892  (*[3], Res*)val (p,_,f,nxt,_,pt) = me nxt'''''_' p'''''_' [] pt'''''_'; (*nxt = Check_elementwise "Assumptions"*)
   893  (*[4], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Check_Postcond ["sqroot-test", "univariate", "equation", "test"]*)
   894  (*[], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = End_Proof'*)
   895 
   896 (*/--- final test ---------------------------------------------------------------------------\\*)
   897 if p = ([], Res) andalso f2str f = "[x = 1]" andalso pr_ctree pr_short pt =
   898    ".    ----- pblobj -----\n" ^
   899    "1.   x + 1 = 2\n" ^
   900    "2.   x + 1 + - 1 * 2 = 0\n" ^
   901    "3.    ----- pblobj -----\n" ^
   902    "3.1.   - 1 + x = 0\n" ^
   903    "3.2.   x = 0 + - 1 * - 1\n" ^
   904    "3.2.1.   x = 0 + - 1 * - 1\n" ^
   905    "3.2.2.   x = 0 + 1\n" (*ATTENTION: see complete Calc below*)
   906 then case nxt of End_Proof' => () | _ => error "re-build: fun locate_input_term CHANGED 1"
   907 else error "re-build: fun locate_input_term CHANGED 2";
   908 
   909 Test_Tool.show_pt_tac pt; (*[
   910 ([], Frm), solve (x + 1 = 2, x)
   911 . . . . . . . . . . Apply_Method ["Test", "squ-equ-test-subpbl1"],
   912 ([1], Frm), x + 1 = 2
   913 . . . . . . . . . . Rewrite_Set "norm_equation",
   914 ([1], Res), x + 1 + - 1 * 2 = 0
   915 . . . . . . . . . . Rewrite_Set "Test_simplify",
   916 ([2], Res), - 1 + x = 0
   917 . . . . . . . . . . Subproblem (Test, ["LINEAR", "univariate", "equation", "test"]),
   918 ([3], Pbl), solve (- 1 + x = 0, x)
   919 . . . . . . . . . . Apply_Method ["Test", "solve_linear"],
   920 ([3,1], Frm), - 1 + x = 0
   921 . . . . . . . . . . Rewrite_Set_Inst ([(''bdv'', x)], "isolate_bdv"),
   922 ([3,1], Res), x = 0 + - 1 * - 1
   923 . . . . . . . . . . Derive Test_simplify,
   924 ([3,2,1], Frm), x = 0 + - 1 * - 1
   925 . . . . . . . . . . Rewrite ("#: - 1 * - 1 = 1", "- 1 * - 1 = 1"),
   926 ([3,2,1], Res), x = 0 + 1
   927 . . . . . . . . . . Rewrite ("radd_0", "0 + ?k = ?k"),
   928 ([3,2,2], Res), x = 1
   929 . . . . . . . . . . Tactic.input_to_string not impl. for ?!,
   930 ([3,2], Res), x = 1
   931 . . . . . . . . . . Check_Postcond ["LINEAR", "univariate", "equation", "test"],
   932 ([3], Res), [x = 1]
   933 . . . . . . . . . . Check_Postcond ["sqroot-test", "univariate", "equation", "test"],
   934 ([], Res), [x = 1]]*)
   935 
   936 \<close> ML \<open>
   937 \<close> ML \<open>
   938 \<close> ML \<open>
   939 \<close> ML \<open>
   940 \<close> ML \<open>
   941 \<close> ML \<open>
   942 \<close> ML \<open>
   943 \<close> ML \<open>
   944 \<close> ML \<open>
   945 \<close>
   946   ML_file "Interpret/step-solve.sml"
   947 
   948 ML_file "MathEngine/me-misc.sml"
   949   ML_file "MathEngine/fetch-tactics.sml"
   950 (** )ML_file "MathEngine/solve.sml"( *loops with eliminate ThmC.numerals_to_Free*)
   951   ML_file "MathEngine/step.sml"
   952 (** )
   953   ML_file "MathEngine/mathengine-stateless.sml"    (*!part. WN130804: +check Interpret/me.sml*)
   954                                    ( *loops with eliminate ThmC.numerals_to_Free*)
   955 ML_file "MathEngine/messages.sml"
   956   ML_file "MathEngine/states.sml"
   957 
   958   ML_file "BridgeLibisabelle/thy-present.sml"
   959   ML_file "BridgeLibisabelle/mathml.sml"           (*part.*)
   960   ML_file "BridgeLibisabelle/pbl-met-hierarchy.sml"
   961   ML_file "BridgeLibisabelle/thy-hierarchy.sml"
   962   ML_file "BridgeLibisabelle/interface-xml.sml"     (*TODO after 2009-2*)
   963 (** )ML_file "BridgeLibisabelle/interface.sml"( *loops with eliminate ThmC.numerals_to_Free*)
   964 (*WITHOUT inhibit exn WN1130621 Isabelle2012-->13 !thehier! THIS ERROR OCCURS:
   965   ... SAME ERROR HERE ON ISABELLE2012 AS IN ISAC ON ISABELLE2011*)
   966 
   967   ML_file "BridgeJEdit/parseC.sml"
   968   ML_file "BridgeJEdit/preliminary.sml"
   969 
   970   ML_file "Knowledge/delete.sml"
   971   ML_file "Knowledge/descript.sml"
   972   ML_file "Knowledge/simplify.sml"
   973   ML_file "Knowledge/poly.sml"
   974   ML_file "Knowledge/gcd_poly_ml.sml"
   975   ML_file "Knowledge/gcd_poly_winkler.sml" (*must be after gcd_poly_ml.sml: redefines functions*)
   976   ML_file "Knowledge/rational.sml"                                            (*Test_Isac_Short*)
   977   ML_file "Knowledge/equation.sml"
   978   ML_file "Knowledge/root.sml"
   979   ML_file "Knowledge/lineq.sml"
   980 
   981 (*ML_file "Knowledge/rooteq.sml"    some complicated equations not recovered from 2002 *)
   982 (*ML_file "Knowledge/rateq.sml"     some complicated equations not recovered----Test_Isac_Short*)
   983   ML_file "Knowledge/r/ootrat.sml"
   984   ML_file "Knowledge/rootrateq.sml"(*ome complicated equations not recovered from 2002 *)
   985 (*ML_file "Knowledge/partial_fractions.sml"  hangs with ML_system_64 = "true"---Test_Isac_Short*)
   986   ML_file "Knowledge/polyeq-1.sml"
   987 (*ML_file "Knowledge/polyeq-2.sml"                                              Test_Isac_Short*)
   988 (*ML_file "Knowledge/rlang.sml"     much to clean up, similar tests in other files     *)
   989   ML_file "Knowledge/calculus.sml"
   990   ML_file "Knowledge/trig.sml"
   991 (*ML_file "Knowledge/logexp.sml"    not included as stuff for presentation of authoring*) 
   992   ML_file "Knowledge/diff.sml"
   993   ML_file "Knowledge/integrate.sml"
   994   ML_file "Knowledge/eqsystem.sml"
   995   ML_file "Knowledge/test.sml"
   996   ML_file "Knowledge/polyminus.sml"
   997   ML_file "Knowledge/vect.sml"
   998   ML_file "Knowledge/diffapp.sml"        (* postponed to dev. specification | TP-prog. *)
   999   ML_file "Knowledge/biegelinie-1.sml"
  1000 (*ML_file "Knowledge/biegelinie-2.sml"                                          Test_Isac_Short*)
  1001 (*ML_file "Knowledge/biegelinie-3.sml"                                          Test_Isac_Short*)
  1002 (** )ML_file "Knowledge/biegelinie-4.sml"( *loops with eliminate ThmC.numerals_to_Free*)
  1003 (** )ML_file "Knowledge/algein.sml"( *loops with eliminate ThmC.numerals_to_Free*)
  1004   ML_file "Knowledge/diophanteq.sml"
  1005 (*ML_file "Knowledge/inverse_z_transform.sml"hangs with ML_system_64 = "true"---Test_Isac_Short*)
  1006   ML_file "Knowledge/inssort.sml"
  1007   ML_file "Knowledge/isac.sml"
  1008   ML_file "Knowledge/build_thydata.sml"
  1009 
  1010   ML_file "Test_Code/test-code.sml"
  1011 
  1012 section \<open>further tests additional to src/.. files\<close>
  1013   ML_file "BridgeLibisabelle/use-cases.sml"
  1014 
  1015   ML \<open>"%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%";\<close>
  1016   ML \<open>"%%%%%%%%%%%%%%%%% end Test_Isac %%%%%%%%%%%%%%%%%%%%%%%%";\<close>
  1017   ML \<open>"%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%";\<close>
  1018 ML \<open>
  1019 \<close> ML \<open>
  1020 \<close> ML \<open>
  1021 \<close> ML \<open>
  1022 \<close> ML \<open>
  1023 \<close> ML \<open>
  1024 \<close> ML \<open>
  1025 \<close> ML \<open>
  1026 \<close> ML \<open>
  1027 \<close> ML \<open>
  1028 \<close> ML \<open>
  1029 \<close>
  1030 
  1031 section \<open>history of tests\<close>
  1032 text \<open>
  1033   Systematic regression tests have been introduced to isac development in 2003.
  1034   Sanity of the regression tests suffers from updates following Isabelle development,
  1035   which mostly exceeded the resources available in isac's development.
  1036 
  1037   The survey below shall support to efficiently use the tests for isac 
  1038   on different Isabelle versions. Conclusion in most cases will be: 
  1039 
  1040   !!! Use most recent tests or go back to the old notebook
  1041       with isac on Isabelle2002. !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
  1042 \<close>
  1043 
  1044 
  1045 subsection \<open>isac on Isabelle2017\<close>
  1046 subsubsection \<open>Summary of development\<close>
  1047 text \<open>
  1048   * Add further signatures, separate structures and cleanup respective files.
  1049   * Show feasibility of moving Isac programs to partial_function, i.e. to the function package.
  1050   * Clean theory dependencies.
  1051   * Start preparing shift from isac-java to Isabelle/jEdit.
  1052 \<close>
  1053 subsubsection \<open>State of tests: unchanged\<close>
  1054 subsubsection \<open>Changesets of begin and end\<close>
  1055 text \<open>
  1056   last changeset with Test_Isac 925fef0f4c81
  1057   first changeset with Test_Isac bbb414976dfe
  1058 \<close>
  1059 
  1060 subsection \<open>isac on Isabelle2015\<close>
  1061 subsubsection \<open>Summary of development\<close>
  1062 text \<open>
  1063   * Add signatures from top of thy-hierarchy down to Interpret (not ProgLang).
  1064     This complicates Test_Isac, see "Prepare running tests" above.
  1065   * Remove TTY interface.
  1066   * Re-activate insertion sort.
  1067 \<close>
  1068 subsubsection \<open>State of tests: unchanged\<close>
  1069 subsubsection \<open>Changesets of begin and end\<close>
  1070 text \<open>
  1071   last changeset with Test_Isac 2f1b2854927a
  1072   first changeset with Test_Isac ???
  1073 \<close>
  1074 
  1075 subsection \<open>isac on Isabelle2014\<close>
  1076 subsubsection \<open>Summary of development\<close>
  1077 text \<open>
  1078   migration from "isabelle tty" --> libisabelle
  1079 \<close>
  1080 
  1081 subsection \<open>isac on Isabelle2013-2\<close>
  1082 subsubsection \<open>Summary of development\<close>
  1083 text \<open>
  1084   reactivated context_thy
  1085 \<close>
  1086 subsubsection \<open>State of tests\<close>
  1087 text \<open>
  1088   TODO
  1089 \<close>
  1090 subsubsection \<open>Changesets of begin and end\<close>
  1091 text \<open>
  1092   TODO
  1093   :
  1094   : isac on Isablle2013-2
  1095   :
  1096   Changeset: 55318 (03826ceb24da) merged
  1097   User: Walther Neuper <neuper@ist.tugraz.at>
  1098   Date: 2013-12-12 14:27:37 +0100 (7 minutes)
  1099 \<close>
  1100 
  1101 subsection \<open>isac on Isabelle2013-1\<close>
  1102 subsubsection \<open>Summary of development\<close>
  1103 text \<open>
  1104   Isabelle2013-1 was replaced within a few weeks due to problems with the document model;
  1105   no significant development steps for ISAC.
  1106 \<close>
  1107 subsubsection \<open>State of tests\<close>
  1108 text \<open>
  1109   See points in subsection "isac on Isabelle2011", "State of tests".
  1110 \<close>
  1111 subsubsection \<open>Changesets of begin and end\<close>
  1112 text \<open>
  1113   Changeset: 55283 (d6e9a34e7142) notes for resuming work on Polynomial.thy
  1114   User: Walther Neuper <neuper@ist.tugraz.at>
  1115   Date: 2013-12-03 18:13:31 +0100 (8 days)
  1116   :
  1117   : isac on Isablle2013-1
  1118   :
  1119   Changeset: 55279 (130688f277ba) Isabelle2013 --> 2013-1: Test_Isac perfect
  1120   User: Walther Neuper <neuper@ist.tugraz.at>
  1121   Date: 2013-11-21 18:12:17 +0100 (2 weeks)
  1122 
  1123 \<close>
  1124 
  1125 subsection \<open>isac on Isabelle2013\<close>
  1126 subsubsection \<open>Summary of development\<close>
  1127 text \<open>
  1128   # Oct.13: replaced "axioms" by "axiomatization"
  1129   # Oct.13: Mathias Lehnfeld started removing Unsynchornized.ref
  1130   # Sep.13: integrated gcd_poly (functional, without Unsychronized.ref) into
  1131     simplification of multivariate rationals (without improving the rulesets involved).
  1132 \<close>
  1133 subsubsection \<open>Run tests\<close>
  1134 text \<open>
  1135   Is standard now; this subsection will be discontinued under Isabelle2013-1
  1136 \<close>
  1137 subsubsection \<open>State of tests\<close>
  1138 text \<open>
  1139   See points in subsection "isac on Isabelle2011", "State of tests".
  1140   # re-activated listC.sml
  1141 \<close>
  1142 subsubsection \<open>Changesets of begin and end\<close>
  1143 text \<open>
  1144   changeset 52174:8b055b17bd84 --- removed all code concerned with "castab = Unsynchronized.ref"
  1145   User: Mathias Lehnfeld <s1210629013@students.fh-hagenberg.at>
  1146   Date: Tue Nov 19 22:23:30 2013 +0000
  1147   :
  1148   : isac on Isablle2013 
  1149   :
  1150   Changeset: 52061 (4ecea2fcdc2c) --- Build_Isac.thy runs on Isabelle2013
  1151   User: Walther Neuper <neuper@ist.tugraz.at>
  1152   Date: 2013-07-15 08:28:50 +0200 (4 weeks)
  1153 \<close>
  1154 
  1155 subsection \<open>isac on Isabelle2012\<close>
  1156 subsubsection \<open>Summary of development\<close>
  1157 text \<open>
  1158   isac on Isabelle2012 is considered just a transitional stage
  1159   within the update from Isabelle2011 to Isabelle2013; thus no further development of isac;
  1160   For considerations on the transition see 
  1161   $ISABELLE_ISAC/Knowledge/Build_Thydata/thy, section "updating isac..".
  1162 \<close>
  1163 subsubsection \<open>Run tests\<close>
  1164 text \<open>
  1165 $ cd /usr/local/isabisac12/
  1166 $ ./bin/isabelle jedit -l HOL src/Tools/isac/Build_Isac.thy
  1167 $ ./bin/isabelle jedit -l Isac test/Tools/isac/Test_Isac.thy
  1168 \<close>
  1169 subsubsection \<open>State of tests\<close>
  1170 text \<open>
  1171   At least the tests from isac on Isabelle2011 run again.
  1172   However, Test_Isac.thy shows erratic behaviour; no errors are obtained when scrolling 
  1173   in parallel with evaluation.
  1174 
  1175   Counting "error in kernel" for Frontend/interface.sml (the tests considered most significant)
  1176   yields 69 hits, some of which were already present before Isabelle2002-->2009-2
  1177   (i.e. on the old notebook from 2002).
  1178 
  1179   Now many tests with (*...=== inhibit exn ...*) give a reason or at least the origin:
  1180   # === inhibit exn WN1130621 Isabelle2012-->13 !thehier! === ...see Build_Thydata.thy
  1181   # === inhibit exn AK110726 === ...reliable work by Alexander Kargl, most likely go back to 2002
  1182   # === inhibit exn WN1130701 broken at Isabelle2002 --> 2009-2 === , most likely go back to 2002
  1183   Reasons for outcommented tests are also found in Test_Isac.thy near the respective file.sml.
  1184 
  1185   Some tests have been re-activated (e.g. error patterns, fill patterns).
  1186 \<close>
  1187 subsubsection \<open>Changesets of begin and end\<close>
  1188 text \<open>
  1189   Changeset: 52051 (35751d90365e) end of improving tests for isac on Isabelle2012
  1190   User: Walther Neuper <neuper@ist.tugraz.at>
  1191   Date: 2013-07-11 16:58:31 +0200 (4 weeks)
  1192   :
  1193   : isac on Isablle2012 
  1194   :
  1195   Changeset: 48757 (74eb3dfc33cc) updated src from Isabelle2011 to Isabelle2012
  1196   User: Walther Neuper <neuper@ist.tugraz.at>
  1197   Date: 2012-09-24 18:35:13 +0200 (8 months)
  1198   ------------------------------------------------------------------------------
  1199   Changeset: 48756 (7443906996a8) merged
  1200   User: Walther Neuper <neuper@ist.tugraz.at>
  1201   Date: 2012-09-24 18:15:49 +0200 (8 months)
  1202 \<close>
  1203 
  1204 subsection \<open>isac on Isabelle2011\<close>
  1205 subsubsection \<open>Summary of development\<close>
  1206 text \<open>
  1207   isac's mathematics engine has been extended by two developments:
  1208   (1) Isabelle's contexts were introduced by Mathias Lehnfeld
  1209   (2) Z_Transform was introduced by Jan Rocnik, which revealed
  1210     further errors introduced by (1).
  1211   (3) "error patterns" were introduced by Gabriella Daroczy
  1212   Regressions tests have been added for all of these.
  1213 \<close>
  1214 subsubsection \<open>Run tests\<close>
  1215 text \<open>
  1216   $ cd /usr/local/isabisac11/
  1217   $ ./bin/isabelle jedit -l HOL src/Tools/isac/Build_Isac.thy
  1218   $ ./bin/isabelle jedit -l Isac test/Tools/isac/Test_Isac.thy
  1219 \<close>
  1220 subsubsection \<open>State of tests\<close>
  1221 text \<open>
  1222   Systematic efforts outcommented less significant tests by (*...=== inhibit exn ...*) 
  1223   and sometimes give reasons for failing tests.
  1224   (*...=== inhibit exn AK...*) was done by Alexander Kargl; this is reliable
  1225   work, some of which couldn't be revised (and renamed) by WN and thus survived some time.
  1226 
  1227   The most signification tests (in particular Frontend/interface.sml) run,
  1228   however, many "error in kernel" are not caught by an exception.
  1229   ------------------------------------------------------------------------------
  1230   After the changeset below Test_Isac worked with check_unsynchronized_ref ():
  1231   ------------------------------------------------------------------------------
  1232   Changeset: 42457 (ca691a84b81a) PROVISIONALLY MADE TESTS RUN with Unsynchronized.ref
  1233   User: Walther Neuper <neuper@ist.tugraz.at>
  1234   Date: 2012-08-06 10:38:11 +0200 (11 months)
  1235 
  1236 
  1237   The list below records TODOs while producing an ISAC kernel for 
  1238   gdaroczy and jrocnik, wich could NOT be done before all tests are RUNNING
  1239   (so to be resumed with Isabelle2013-1):
  1240   ############## WNxxxxxx.TODO can be found in sources ##############
  1241   --------------------------------------------------------------------------------
  1242   WN111013.TODO: lots of cleanup/removal in test/../Test.thy
  1243   --------------------------------------------------------------------------------
  1244   WN111013.TODO: remove concept around "fun implicit_take", lots of troubles with 
  1245   this special case (see) --- why not nxt = Model_Problem here ? ---
  1246   --------------------------------------------------------------------------------
  1247   WN111014.TODO calculate_Poly < calculate_Rational < calculate_RootRat, see test/
  1248   ... FIRST redesign 
  1249   # simplify_* , *_simp_* 
  1250   # norm_* 
  1251   # calc_* , calculate_*  ... require iteration over all rls ...
  1252   ... see --- val rls = calculate_RootRat > calculate_Rational --- CONTINUE !
  1253   --------------------------------------------------------------------------------
  1254   WN111014.TODO fun prep_rls | !!!use this function in ruleset' := !!!
  1255   --------------------------------------------------------------------------------
  1256   WN120314 changeset a393bb9f5e9f drops root equations.
  1257   see test/Tools/isac/Knowledge/rootrateq.sml 
  1258   --------------------------------------------------------------------------------
  1259   WN120317.TODO changeset 977788dfed26 dropped rateq:
  1260   # test --- repair NO asms from rls RatEq_eliminate --- shows error from 2002
  1261   # test --- solve (1/x = 5, x) by me --- and --- x / (x  \<up>  2 - 6 * x + 9) - ...:    
  1262     investigation Check_elementwise stopped due to too much effort finding out,
  1263     why Check_elementwise worked in 2002 in spite of the error.
  1264   --------------------------------------------------------------------------------
  1265   WN120317.TODO postponed test/../ratinal,ratinal2.sml to joint work with dmeindl 
  1266   --------------------------------------------------------------------------------
  1267   WN120317.TODO found by test --- interSteps for Schalk 299a --- that 
  1268     NO test with 'interSteps' is checked properly (with exn on changed behaviour)
  1269   --------------------------------------------------------------------------------
  1270   WN120317.TODO test --- Matthias Goldgruber 2003 rewrite orders --- has
  1271     a newly outcommented test where rewrite_set_ make_polynomial --> NONE
  1272   --------------------------------------------------------------------------------
  1273   WN120320.TODO check-improve rlsthmsNOTisac:
  1274   DONE make test --- old compute rlsthmsNOTisac by eq_thmI'
  1275   DONE compare rlsthmsNOTisac in thms-survey-Isa02-Isa09-2.sml .. Isac.thy 
  1276   FOUND 120321: Theory.axioms_of doesnt find LENGTH_CONS etc, thus are in Isab
  1277   # mark twice thms (in isac + (later) in Isabelle) in Isac.thy
  1278   --------------------------------------------------------------------------------
  1279   WN120320.TODO rlsthmsNOTisac: replace twice thms  \<up> 
  1280   --------------------------------------------------------------------------------
  1281   WN120320.TODO rlsthmsNOTisac: reconsider design of sym_* thms, see test
  1282   --- OLD compute rlsthmsNOTisac by eq_thmID ---: some are in isab, some in isac.
  1283   --------------------------------------------------------------------------------
  1284   WN120321.TODO rearrange theories:
  1285     Knowledge
  1286       :
  1287       Prog_Expr.thy
  1288       ///Input_Descript.thy --> ProgLang
  1289       Delete.thy   <--- first_Knowledge_thy (*mv to Prog_Expr.thy*)
  1290     ProgLang: restructure Build_Isac.thy such that no xmlsrc.thy, ProgLang.thy
  1291           Interpret.thy are generated (simplifies xml structure for theories)
  1292       Program.thy
  1293       Tools.thy
  1294       ListC.thy    <--- first_Proglang_thy
  1295   --------------------------------------------------------------------------------
  1296   WN120321.TODO reanimate test/../simplify.sml: CAS-command Simplify
  1297       EXN "simplify.sml: diff.behav. CAScmd: Simplify (2*a + 3*a)"
  1298   broken during work on thy-hierarchy
  1299   --------------------------------------------------------------------------------
  1300   WN120321.TODO LAST in IsacScripts + in IsacKnowledge, see
  1301   test --- the_hier (get_thes ()) (collect_thydata ())---
  1302   --------------------------------------------------------------------------------
  1303   WN120405a.TODO src/../pbl-met-hierarchy.sml: fun pbl2term--> Isac' instead Isac_
  1304   !!add mutual crossreferences to ?fun headline??? where the same has to be done:
  1305   !!OR BETTER: use only 2 functions for adding/removing "'" to/from thy!!
  1306   --------------------------------------------------------------------------------
  1307   WN120411 scanning html representation of newly generated knowledge:
  1308   * thy:
  1309   ** Theorems: only "Proof of the theorem" (correct!)
  1310                and "(c) isac-team (math-autor)"
  1311   ** Rulesets: only "Identifier:///"
  1312                and "(c) isac-team (math-autor)"
  1313   ** IsacKnowledge: link to dependency graph (which needs to be created first)
  1314   ** IsacScripts --> ProgramLanguage
  1315   *** Tools: Theorems: arity_type_cp, arity_type_nam, ... wegnehmen
  1316   
  1317   * pbl: OK !?!
  1318   * met: OK !?!
  1319   * exp: 
  1320   ** Z-Transform is missing !!!
  1321   ** type-constraints !!!
  1322   --------------------------------------------------------------------------------
  1323   WN120417: merging xmldata revealed:
  1324   ..............NEWLY generated:........................................
  1325   <THEOREMDATA>
  1326     <GUH> thy_isab_Fun-thm-o_apply </GUH>
  1327     <STRINGLIST>
  1328       <STRING> Isabelle </STRING>
  1329       <STRING> Fun </STRING>
  1330       <STRING> Theorems </STRING>
  1331       <STRING> o_apply </STRING>
  1332     </STRINGLIST>
  1333       <MATHML>
  1334         <ISA> (?f o ?g) ?x = ?f (?g ?x) </ISA>
  1335       </MATHML>  <PROOF>
  1336       <EXTREF>
  1337         <TEXT> Proof of the theorem </TEXT>
  1338         <URL> http://www.ist.tugraz.at/projects/isac/www/kbase/thy/browser_info/HOL/HOL-Real/Isac/Fun.html </URL>
  1339       </EXTREF>
  1340     </PROOF>
  1341     <EXPLANATIONS> </EXPLANATIONS>
  1342     <MATHAUTHORS>
  1343       <STRING> Isabelle team, TU Munich </STRING>
  1344     </MATHAUTHORS>
  1345     <COURSEDESIGNS>
  1346     </COURSEDESIGNS>
  1347   </THEOREMDATA>
  1348   ..............OLD FORMAT:.............................................
  1349   <THEOREMDATA>
  1350     <GUH> thy_isab_Fun-thm-o_apply </GUH>
  1351     <STRINGLIST>
  1352       <STRING> Isabelle </STRING>
  1353       <STRING> Fun </STRING>
  1354       <STRING> Theorems </STRING>
  1355       <STRING> o_apply </STRING>
  1356     </STRINGLIST>
  1357     <THEOREM>
  1358       <ID> o_apply </ID>
  1359       <MATHML>
  1360         <ISA> (?f o ?g) ?x = ?f (?g ?x) </ISA>
  1361       </MATHML>
  1362     </THEOREM>
  1363     <PROOF>
  1364       <EXTREF>
  1365         <TEXT> Proof of the theorem </TEXT>
  1366         <URL> http://www.ist.tugraz.at/projects/isac/www/kbase/thy/browser_info/HOL/HOL-Real/Isac/Fun.html </URL>
  1367       </EXTREF>
  1368     </PROOF>
  1369     <EXPLANATIONS> </EXPLANATIONS>
  1370     <MATHAUTHORS>
  1371       <STRING> Isabelle team, TU Munich </STRING>
  1372     </MATHAUTHORS>
  1373     <COURSEDESIGNS>
  1374     </COURSEDESIGNS>
  1375   </THEOREMDATA>
  1376   --------------------------------------------------------------------------------
  1377 \<close>
  1378 subsubsection \<open>Changesets of begin and end\<close>
  1379 text \<open>
  1380   isac development was done between these changesets:
  1381   ------------------------------------------------------------------------------
  1382   Changeset: 42519 (1f3b4270363e) meeting dmeindl: added missing files
  1383   User: Walther Neuper <neuper@ist.tugraz.at>
  1384   Date: 2012-09-24 16:39:30 +0200 (8 months)
  1385   :
  1386   : isac on Isablle2011
  1387   :
  1388   Changeset:41897 (355be7f60389) merged isabisac with Isabelle2011
  1389   Branch: decompose-isar 
  1390   User: Walther Neuper <neuper@ist.tugraz.at>
  1391   Date: 2011-02-25 13:04:56 +0100 (2011-02-25)
  1392   ------------------------------------------------------------------------------
  1393 \<close>
  1394 
  1395 subsection \<open>isac on Isabelle2009-2\<close>
  1396 subsubsection \<open>Summary of development\<close>
  1397 text \<open>
  1398   In 2009 the update of isac from Isabelle2002 started with switching from CVS to hg.
  1399   The update was painful (bridging 7 years of Isabelle development) and cut short 
  1400   due to the start of introducing Isabelle's contexts (Mathias Lehnfeld) and
  1401   going on to Isabelle2011 although most of the tests did not run.
  1402 \<close>
  1403 subsubsection \<open>Run tests\<close>
  1404 text \<open>
  1405   WN131021 this is broken by installation of Isabelle2011/12/13,
  1406   because all these write their binaries to ~/.isabelle/heaps/..
  1407 
  1408   $ cd /usr/local/isabisac09-2/
  1409   $ ./bin/isabelle emacs -l HOL src/Tools/isac/Build_Isac.thy
  1410   $ ./bin/isabelle emacs -l Isac src/Tools/isac/Test_Isac.thy
  1411   NOT THE RIGHT VERSION.....    test/Tools/isac/Test_Isac.thy !!!
  1412 \<close>
  1413 subsubsection \<open>State of tests\<close>
  1414 text \<open>
  1415   Most tests are broken by the update from Isabelle2002 to Isabelle2009-2.
  1416 \<close>
  1417 subsubsection \<open>Changesets of begin and end\<close>
  1418 text \<open>
  1419   isac development was done between these changesets:
  1420   ------------------------------------------------------------------------------
  1421   Changeset: 38115 (940a5feea094) Little improvements of isac-plugin
  1422   Branch: decompose-isar 
  1423   User: Marco Steger <m.steger@student.tugraz.at>
  1424   Date: 2011-02-06 18:30:28 +0100 (2011-02-06)
  1425   :
  1426   : isac on Isablle2009-2
  1427   :
  1428   Changeset: 37870 (5100a9c3abf8) created branch isac-from-Isabelle2009-2
  1429   Branch: isac-from-Isabelle2009-2 
  1430   User: Walther Neuper <neuper@ist.tugraz.at>
  1431   Date: 2010-07-21 09:59:35 +0200 (2010-07-21)
  1432   ------------------------------------------------------------------------------
  1433 \<close>
  1434 
  1435 subsection \<open>isac on Isabelle2002\<close>
  1436 subsubsection \<open>Summary of development\<close>
  1437 text \<open>
  1438   From 1999 to 2010 all the basic functionality (except "ctxt" and "error pattern")
  1439   of isac's mathematics engine has been implemented.
  1440 \<close>
  1441 subsubsection \<open>Run tests\<close>
  1442 subsubsection \<open>State of tests\<close>
  1443 text \<open>
  1444   All tests work on an old notebook (the right PolyML coudn't be upgraded to more
  1445   recent Linux versions)
  1446 \<close>
  1447 subsubsection \<open>Changesets of begin and end\<close>
  1448 text \<open>
  1449   Up to the above Mercurial changeset 5100a9c3abf8 isac used CVS;
  1450   see https://intra.ist.tugraz.at/hg/isac containing a conversion to Mercurial.
  1451 \<close>
  1452 
  1453 end
  1454 (*========== inhibit exn 130719 Isabelle2013 ===================================
  1455 ============ inhibit exn 130719 Isabelle2013 =================================*)
  1456 
  1457 (*-.-.-.-.-.-isolate response.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.
  1458   -.-.-.-.-.-.-isolate response.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.*)
  1459