test/Tools/isac/Test_Isac_Short.thy
author wneuper <walther.neuper@jku.at>
Thu, 15 Jul 2021 20:02:16 +0200
changeset 60325 a7c0e6ab4883
parent 60324 5c7128feb370
child 60326 33e04eb1a2f0
permissions -rw-r--r--
rewrite.sml + poly.sml + rational.sml + polyminus.sml: ok
     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 (*+*)val Test_Out.FormKF "- 9 + 12 + 5 * e - 7 * e + 6 * f - 4 * f - 8 * g + 10 * g" = f
   521 \<close> ML \<open>
   522 \<close> ML \<open> (*GOON*)
   523 map Tactic.input_to_string (specific_from_prog pt p)
   524 \<close> ML \<open>
   525 (*+*)if map Tactic.input_to_string (specific_from_prog pt p) =
   526    ["Rewrite (\"subtrahiere_x_plus_minus\", \"\<lbrakk>?l is_const; ?m is_const\<rbrakk>\n\<Longrightarrow> ?x + ?m * ?v - ?l * ?v = ?x + (?m - ?l) * ?v\")",
   527     "Rewrite (\"subtrahiere_x_minus_plus\", \"\<lbrakk>?l is_const; ?m is_const\<rbrakk>\n\<Longrightarrow> ?x - ?m * ?v + ?l * ?v = ?x + (- ?m + ?l) * ?v\")"]
   528   then () else error "specific_from_prog ([1], Res) CHANGED";
   529 (*[2], Res*)val ("ok", (_, _, ptp as (pt, p))) = Step.by_tactic (hd (specific_from_prog pt p)) (pt, p);
   530 
   531 \<close> ML \<open>
   532 (*+*)if map Tactic.input_to_string (specific_from_prog pt p) =
   533    ["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\")",
   534     "Rewrite (\"subtrahiere_x_plus_minus\", \"\<lbrakk>?l is_const; ?m is_const\<rbrakk>\n\<Longrightarrow> ?x + ?m * ?v - ?l * ?v = ?x + (?m - ?l) * ?v\")",
   535     "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"]
   536   then () else error "specific_from_prog ([1], Res) CHANGED";
   537 (* = ([3], Res)*)val ("ok", (_, _, ptp as (pt, p))) = Step.by_tactic (hd (specific_from_prog pt p)) (pt, p);
   538 
   539 \<close> ML \<open>
   540 (*//----------------- exception PTREE "get_obj f EmptyPtree" raised --------------------------\\*)
   541 (** )val ("ok", (_, _, ptp as (pt, p))) =( **)
   542       Step.by_tactic (hd (specific_from_prog pt p)) (pt, p);
   543 "~~~~~ fun by_tactic , args:"; val (tac, (ptp as (pt, p))) = (hd (specific_from_prog pt p), (pt, p));
   544       val Applicable.Yes m = (*case*) Solve_Step.check tac (pt, p) (*of*);
   545       (*if*) Tactic.for_specify' m; (*false*)
   546 
   547 Step_Solve.by_tactic m (pt, p);
   548 "~~~~~ fun by_tactic , args:"; val (m, (pt, po as (p, p_))) = (m, (pt, p));
   549     (*if*) MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p) (*else*);
   550 	      val thy' = get_obj g_domID pt (par_pblobj pt p);
   551 	      val (is, sc) = LItool.resume_prog thy' (p,p_) pt;
   552 
   553   (*case*) locate_input_tactic sc (pt, po) (fst is) (snd is) m (*of*);
   554 "~~~~~ fun locate_input_tactic , args:"; val ((Rule.Prog prog), (cstate as (pt, (*?*)pos(*?*))), istate, ctxt, tac)
   555   = (sc, (pt, po), (fst is), (snd is), m);
   556       val srls = LItool.get_simplifier cstate (*TODO: shift into Istate.T*);
   557 
   558   (*case*) scan_to_tactic1 (prog, (cstate, ctxt, tac)) istate (*of*);
   559 "~~~~~ fun scan_to_tactic1 , args:"; val ((prog, (cctt as ((_, p), _, _))), (Istate.Pstate (ist as {path, ...})))
   560   = ((prog, (cstate, ctxt, tac)), istate);
   561     (*if*) path = [] orelse ((last_elem o fst) p = 0 andalso snd p = Res) (*else*);
   562 
   563            go_scan_up1 (prog, cctt) ist;
   564 "~~~~~ fun go_scan_up1 , args:"; val ((pcct as (prog, _)), (ist as {path, ...}))
   565   = ((prog, cctt), ist);
   566   (*if*) 1 < length path (*then*);
   567 
   568            scan_up1 pcct (ist |> path_up) (TermC.sub_at (path_up' path) prog);
   569 "~~~~~ and scan_up1 , args:"; val (pcct, ist, (Const ("Tactical.Try"(*2*), _) $ _))
   570   = (pcct, (ist |> path_up), (TermC.sub_at (path_up' path) prog));
   571 
   572            go_scan_up1 pcct ist;
   573 "~~~~~ and go_scan_up1 , args:"; val ((pcct as (prog, _)), (ist as {path, ...}))
   574   = (pcct, ist);
   575   (*if*) 1 < length path (*then*);
   576 
   577            scan_up1 pcct (ist |> path_up) (TermC.sub_at (path_up' path) prog);
   578 "~~~~~ and scan_up1 , args:"; val ((pcct as (prog, cct as (cstate, _, _))), ist,
   579     (Const ("Tactical.Chain"(*3*), _) $ _ ))
   580   = (pcct, (ist |> path_up), (TermC.sub_at (path_up' path) prog));
   581        val e2 = check_Seq_up ist prog
   582 ;
   583   (*case*) scan_dn1 cct (ist |> path_up_down [R] |> set_or ORundef) e2 (*of*);
   584 "~~~~~ fun scan_dn1 , args:"; val (yyy, ist, (Const ("Tactical.Chain"(*2*), _) $ e1 $ e2))
   585   = (cct, (ist |> path_up_down [R] |> set_or ORundef), e2);
   586 
   587   (*case*) scan_dn1 cct (ist |> path_down [L, R]) e1 (*of*);
   588 "~~~~~ fun scan_dn1 , args:"; val (xxx, ist, (Const ("Tactical.Try"(*2*), _) $ e))
   589   = (cct, (ist |> path_down [L, R]), e1);
   590 
   591   (*case*) scan_dn1 cct (ist |> path_down [R]) e (*of*);
   592     (*======= end of scanning tacticals, a leaf =======*)
   593 "~~~~~ fun scan_dn1 , args:"; val ((cct as (_, ctxt, _)), (ist as {eval, ...}), t)
   594   = (cct, (ist |> path_down [R]), e);
   595     (*if*) Tactical.contained_in t (*else*);
   596   val (Program.Tac prog_tac, form_arg) = (*case*)
   597     LItool.check_leaf "locate" ctxt eval (get_subst ist) t (*of*);
   598 
   599            check_tac1 cct ist (prog_tac, form_arg);
   600 "~~~~~ fun check_tac1 , args:"; val (((pt, p), ctxt, tac), (ist as {act_arg, or, ...}), (prog_tac, form_arg)) =
   601   (cct, ist, (prog_tac, form_arg));
   602 val LItool.Not_Associated = (*case*)
   603   LItool.associate pt ctxt (tac, prog_tac) (*of*);
   604      val _(*ORundef*) = (*case*) or (*of*);
   605      val Applicable.Yes m' =
   606           (*case*) Solve_Step.check (LItool.tac_from_prog pt (ThyC.get_theory "Isac_Knowledge") prog_tac) (pt, p) (*of*);
   607 
   608   Reject_Tac1 (ist |> set_subst_false (form_arg, Tactic.result m'), ctxt, tac)
   609           (*return from check_tac1*);
   610 "~~~~~ from fun check_tac1 \<longrightarrow>fun scan_dn1 \<longrightarrow>fun scan_dn1  \<longrightarrow>fun locate_input_tactic , return:"; val (Reject_Tac1 _) =
   611   (Reject_Tac1 (ist |> set_subst_false (form_arg, Tactic.result m'), ctxt, tac));
   612 
   613 (*/----- original before child of 7e314dd233fd -------------------------------------------------\* )
   614     val (Program.Tac prog_tac, form_arg) = (*case*) check_leaf "locate" ctxt eval (get_subst ist) t (*of*);
   615       val Not_Associated = (*case*) associate pt ctxt (tac, stac) (*of*);
   616       val ORundef = (*case*) or (*of*);
   617   val Applicable.No "norm_equation not applicable" =    
   618       (*case*) Solve_Step.check (LItool.tac_from_prog pt (ThyC.get_theory "Isac_Knowledge") stac) (pt, p) (*of*);
   619 
   620    (Term_Val1 act_arg) (* return value *);
   621 
   622 val Rewrite' ("PolyMinus", "tless_true", _, _, ("tausche_minus",_ (*"?b ist_monom \<Longrightarrow> ?a kleiner ?b \<Longrightarrow> ?b - ?a = - ?a + ?b"*)),
   623    t, (res, asm)) = m;
   624 
   625 if pstate2str ist =
   626   "([\"\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" ^
   627   "- (8 * g) + (- 9 + 12 + 5 * e - 7 * e + (6 - 4) * f) + 10 * g, ORundef, true, false)"
   628 andalso
   629   UnparseC.term t = "- (8 * g) + (- 9 + 12 + 5 * e - 7 * e + (6 - 4) * f) + 10 * g"
   630 andalso
   631   UnparseC.term res = "- (8 * g) + (- 9 + 12 + 5 * e - 7 * e + (- 4 + 6) * f) + 10 * g"
   632 andalso
   633   UnparseC.terms asm = "[\"4 kleiner 6\",\"6 ist_monom\"]"
   634 then () else error "locate_input_tactic Helpless, but applicable CHANGED";
   635 ( *\----- original before child of 7e314dd233fd -------------------------------------------------/*)
   636 
   637 
   638 \<close> ML \<open>
   639 "----------- re-build: fun find_next_step, mini ------------------------------------------------";
   640 "----------- re-build: fun find_next_step, mini ------------------------------------------------";
   641 "----------- re-build: fun find_next_step, mini ------------------------------------------------";
   642 val fmz = ["Term (a + a ::real)", "normalform n_n"];
   643 val (dI',pI',mI') = ("Poly",["polynomial", "simplification"],["simplification", "for_polynomials"]);
   644 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   645 (*[], Pbl*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, Pos.e_pos'), []);(*Model_Problem*)
   646 (*[], Pbl*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, Pos.e_pos'), []);(*Specify_Theory "Poly"*)
   647 (*[], Pbl*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, Pos.e_pos'), []);(*Specify_Problem ["polynomial", "simplification"]*)
   648 (*[], Met*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, e_pos'), []);(*Specify_Method  ["simplification", "for_polynomials"]*)
   649 (*1], Frm*)val (_, ([(tac, _, _)], _, (pt, p))) = Step.do_next p ((pt, e_pos'), []);(*Apply_Method ["simplification", "for_polynomials"]*)
   650 (*[1], Res*)val (_, ([(tac'''''_', _, _)], _, (pt'''''_', p'''''_'))) =
   651 
   652       Step.do_next p ((pt, e_pos'), []);(*Rewrite_Set "norm_Poly"*)
   653 \<close> ML \<open>
   654 (*//------------------ go into 1 ------------------------------------------------------------\\*)
   655 "~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis))
   656   = (p, ((pt, e_pos'), []));
   657   val pIopt = Ctree.get_pblID (pt, ip);
   658     (*if*)  ip = ([], Res) (*else*);
   659       val _ = (*case*) tacis (*of*);
   660       val SOME _ = (*case*) pIopt (*of*);
   661       (*if*) member op = [Pos.Pbl, Pos.Met] p_ (*else*);
   662 
   663 val ("ok", ([(Rewrite_Set "norm_Poly", _, _)], _, (_, ([1], Res)))) =
   664 Step_Solve.do_next (pt, ip);
   665 "~~~~~ and do_next , args:"; val (ptp as (pt, pos as (p, p_))) =  (pt, ip);
   666     (*if*) MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p) (*else*);
   667         val thy' = get_obj g_domID pt (par_pblobj pt p);
   668 	      val ((ist, ctxt), sc) = LItool.resume_prog thy' (p,p_) pt;
   669 
   670 val Next_Step (_, _, Rewrite_Set' ("Poly", _, Rule_Set.Sequence {id = "norm_Poly", ...}, _, _)) =
   671            LI.find_next_step sc (pt, pos) ist ctxt (*of*);
   672 "~~~~~ fun find_next_step , args:"; val ((Rule.Prog prog), (ptp as(pt, (p, _))), (Pstate ist), ctxt)
   673   = (sc, (pt, pos), ist, ctxt);
   674 
   675 val Accept_Tac (Rewrite_Set' ("Poly", _, Rule_Set.Sequence {id = "norm_Poly", ...}, _, _), _, _) =
   676   (*case*) scan_to_tactic (prog, (ptp, ctxt)) (Pstate ist) (*of*);
   677 "~~~~~ fun scan_to_tactic , args:"; val ((prog, cc), (Pstate (ist as {path, ...})))
   678   = ((prog, (ptp, ctxt)), (Pstate ist));
   679   (*if*) path = [] (*then*);
   680 
   681 val Accept_Tac (Rewrite_Set' ("Poly", _, Rule_Set.Sequence {id = "norm_Poly", ...}, _, _), _, _) =
   682             scan_dn cc (trans_scan_dn ist) (Program.body_of prog);
   683 "~~~~~ fun scan_dn , args:"; val ((cc as (_, ctxt)), (ist as {eval, ...}), t)
   684   = (cc, (trans_scan_dn ist), (Program.body_of prog));
   685     (*if*) Tactical.contained_in t (*else*);
   686       val (Program.Tac prog_tac, form_arg) = (*case*) LItool.check_leaf "next  " ctxt eval (get_subst ist) t (*of*);
   687 
   688 val Accept_Tac (Rewrite_Set' ("Poly", _, Rule_Set.Sequence {id = "norm_Poly", ...}, _, _), _, _) =
   689           check_tac cc ist (prog_tac, form_arg)  (*return from xxx*);
   690 "~~~~~ from fun scan_dn\<longrightarrow>fun scan_to_tactic\<longrightarrow>fun find_next_step, return:"; val (Accept_Tac (tac, ist, ctxt))
   691   = (check_tac cc ist (prog_tac, form_arg));
   692 
   693     Next_Step (Pstate ist, Tactic.insert_assumptions tac ctxt, tac)  (*return from find_next_step*);
   694 "~~~~~ from fun find_next_step\<longrightarrow>and do_next\<longrightarrow>fun zzz, return:"; val (Next_Step (ist, ctxt, tac))
   695   = (Next_Step (Pstate ist, Tactic.insert_assumptions tac ctxt, tac));
   696 
   697            LI.by_tactic tac (ist, Tactic.insert_assumptions tac ctxt) ptp  (*return from and do_next*);
   698 "~~~~~ from and do_next\<longrightarrow>fun do_next\<longrightarrow>toplevel, return:"; val  (_, ([(tac''''', _, _)], _, (pt''''', p''''')))
   699   = (LI.by_tactic tac (ist, Tactic.insert_assumptions tac ctxt) ptp);
   700 (*\\------------------ end of go into 1 -----------------------------------------------------//*)
   701 
   702 (*[], Res*)val (_, ([(tac''''', _, _)], _, (pt''''', p'''''))) =
   703 
   704       Step.do_next p'''''_' ((pt'''''_', Pos.e_pos'), []);(* Check_Postcond ["polynomial", "simplification"]*)
   705 (*//------------------ go into 2 ------------------------------------------------------------\\*)
   706 "~~~~~ fun do_next , args:"; val ((ip as (_, p_)), (ptp as (pt, p), tacis))
   707   = (p''''', ((pt''''', e_pos'), []));
   708   val pIopt = Ctree.get_pblID (pt, ip);
   709     (*if*)  ip = ([], Res) (*else*);
   710       val _ = (*case*) tacis (*of*);
   711       val SOME _ = (*case*) pIopt (*of*);
   712       (*if*) member op = [Pos.Pbl, Pos.Met] p_ (*else*);
   713 
   714 val ("ok", ([(Check_Postcond ["polynomial", "simplification"], _, _)], _, (_, ([], Res)))) =
   715 Step_Solve.do_next (pt, ip);
   716 "~~~~~ and do_next , args:"; val (ptp as (pt, pos as (p, p_))) =  (pt, ip);
   717     (*if*) MethodC.id_empty = get_obj g_metID pt (par_pblobj pt p) (*else*);
   718         val thy' = get_obj g_domID pt (par_pblobj pt p);
   719 	      val ((ist, ctxt), sc) = LItool.resume_prog thy' (p,p_) pt;
   720 
   721   (** )val End_Program (ist, tac) = 
   722  ( *case*) LI.find_next_step sc (pt, pos) ist ctxt (*of*);
   723 "~~~~~ fun find_next_step , args:"; val ((Rule.Prog prog), (ptp as(pt, (p, _))), (Pstate ist), ctxt)
   724   = (sc, (pt, pos), ist, ctxt);
   725 
   726 (*  val Term_Val (Const ("Groups.times_class.times", _) $ Free ("2", _) $ Free ("a", _))*)
   727   (** )val Term_Val prog_result =
   728  ( *case*) scan_to_tactic (prog, (ptp, ctxt)) (Pstate ist) (*of*);
   729 "~~~~~ fun scan_to_tactic , args:"; val ((prog, cc), (Pstate (ist as {path, ...})))
   730   = ((prog, (ptp, ctxt)), (Pstate ist));
   731   (*if*) path = [] (*else*);
   732 
   733            go_scan_up (prog, cc) (trans_scan_up ist |> set_found);
   734 "~~~~~ fun go_scan_up , args:"; val ((pcc as (sc, _)), (ist as {path, act_arg, found_accept, ...}))
   735   = ((prog, cc), (trans_scan_up ist(*|> set_found !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!! *)));
   736     (*if*) path = [R] (*then*);
   737       (*if*) found_accept = true (*then*);
   738 
   739       Term_Val act_arg (*return from go_scan_up*);
   740 "~~~~~ from fun go_scan_up\<longrightarrow>fun scan_to_tactic, return:"; val (Term_Val prog_result) = (Term_Val act_arg);
   741 
   742     Term_Val prog_result  (*return from scan_to_tactic*);
   743 "~~~~~ from fun scan_to_tactic\<longrightarrow>fun find_next_step, return:"; val (Term_Val prog_result) = (Term_Val prog_result);
   744     val (true, p', _) = (*case*) parent_node pt p (*of*);
   745               val (_, pblID, _) = get_obj g_spec pt p';
   746 
   747      End_Program (Pstate ist, Tactic.Check_Postcond' (pblID, prog_result))
   748      (*return from find_next_step*);
   749 "~~~~~ from fun find_next_step\<longrightarrow>and do_next\<longrightarrow>fun zzz, return:"; val (End_Program (ist, tac))
   750   = (End_Program (Pstate ist, Tactic.Check_Postcond' (pblID,prog_result)));
   751       val _ = (*case*) tac (*of*);
   752 
   753 val ("ok", ([(Check_Postcond ["polynomial", "simplification"], _, _)], _, (_, ([], Res))))
   754    = LI.by_tactic tac (ist, ctxt) ptp (*return from and do_next*);
   755 "~~~~~ from and do_next\<longrightarrow>top level, return:"; val (_, ([(tac''''', _, _)], _, (pt''''', p''''')))
   756   = (LI.by_tactic tac (ist, ctxt) ptp);
   757 (*\\------------------ end of go into 2 -----------------------------------------------------//*)
   758 
   759 (*[], Und*)val (msg, ([], _, (pt, p))) = Step.do_next p''''' ((pt''''', Pos.e_pos'), []);(**)
   760 
   761 Test_Tool.show_pt_tac pt; (*[
   762 ([], Frm), Simplify (a + a)
   763 . . . . . . . . . . Apply_Method ["simplification", "for_polynomials"],
   764 ([1], Frm), a + a
   765 . . . . . . . . . . Rewrite_Set "norm_Poly",
   766 ([1], Res), 2 * a
   767 . . . . . . . . . . Check_Postcond ["polynomial", "simplification"],
   768 ([], Res), 2 * a]*)
   769 
   770 (*/--- final test ---------------------------------------------------------------------------\\*)
   771 val (res, asm) = (get_obj g_result pt (fst p));
   772 if UnparseC.term res = "2 * a" andalso map UnparseC.term asm = []
   773 andalso p = ([], Und) andalso msg = "end-of-calculation"
   774 andalso pr_ctree pr_short pt = ".    ----- pblobj -----\n1.   a + a\n"
   775 then 
   776   case tac''''' of Check_Postcond ["polynomial", "simplification"] => () 
   777   | _ => error "re-build: fun find_next_step, mini 1"
   778 else error "re-build: fun find_next_step, mini 2"
   779 
   780 
   781 \<close> ML \<open>
   782 "----------- re-build: fun locate_input_term ---------------------------------------------------";
   783 "----------- re-build: fun locate_input_term ---------------------------------------------------";
   784 "----------- re-build: fun locate_input_term ---------------------------------------------------";
   785 (*cp from inform.sml
   786  ----------- appendFormula: on Res + late deriv ------------------------------------------------*)
   787 val fmz = ["equality (x+1=(2::real))", "solveFor x", "solutions L"];
   788 val (dI',pI',mI') = ("Test", ["sqroot-test", "univariate", "equation", "test"],
   789    ["Test", "squ-equ-test-subpbl1"]);
   790 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   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 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   797 (*[], Met*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Apply_Method ["Test", "squ-equ-test-subpbl1"]*);
   798 
   799 (*[1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p [] pt;(*Rewrite_Set "norm_equation"*)
   800 (*+*)if f2str f = "x + 1 = 2" then () else error "locate_input_term at ([1], Frm) CHANGED";
   801 
   802 (*[1], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt;(*Rewrite_Set "Test_simplify"*)
   803 (*+*)if f2str f = "x + 1 + - 1 * 2 = 0" then () else error "locate_input_term at ([1], Frm) CHANGED";
   804 
   805 Test_Tool.show_pt_tac pt; (*[
   806 ([], Frm), solve (x + 1 = 2, x)
   807 . . . . . . . . . . Apply_Method ["Test", "squ-equ-test-subpbl1"],
   808 ([1], Frm), x + 1 = 2
   809 . . . . . . . . . . Rewrite_Set "norm_equation",
   810 ([1], Res), x + 1 + - 1 * 2 = 0             ///Check_Postcond..ERROR*)
   811 
   812 (*//---------- appendFormula 1 "x = 1" \<longrightarrow> Step_Solve.inform \<longrightarrow> LI.locate_input_term ----------\\*)
   813 "~~~~~ fun appendFormula , args:"; val ((*cI, *) ifo: TermC.as_string) = ((**) "x = 1");
   814     val cs = (*get_calc cI*) ((pt, p), [(*nxt, nxt_, (pos, (ist, ctxt))*)])
   815     val pos = (*get_pos cI 1*) p
   816 
   817 (*+*)val ptp''''' = (pt, p);
   818 (*+*)if snd ptp''''' = ([1], Res) then () else error "old_cs changed";
   819 (*+*)Test_Tool.show_pt_tac pt; (*[
   820 (*+*)([], Frm), solve (x + 1 = 2, x)
   821 (*+*). . . . . . . . . . Apply_Method ["Test", "squ-equ-test-subpbl1"],
   822 (*+*)([1], Frm), x + 1 = 2
   823 (*+*). . . . . . . . . . Rewrite_Set "norm_equation",
   824 (*+*)([1], Res), x + 1 + - 1 * 2 = 0      ///Check_Postcond*)
   825 
   826   val ("ok", cs' as (_, _, ptp')) =
   827     (*case*) Step.do_next pos cs (*of*);
   828 
   829 val ("ok", (_(*use in DG !!!*), [], ptp''''' as (pt''''', p'''''))) = (*case*)
   830      Step_Solve.by_term ptp' (encode ifo) (*of*);
   831 "~~~~~ fun Step_Solve.by_term , args:"; val ((pt, pos as (p, _)), istr)
   832   = (ptp', (encode ifo));
   833   val SOME f_in =
   834     (*case*) TermC.parse (ThyC.get_theory "Isac_Knowledge") istr (*of*);
   835   	  val f_in = Thm.term_of f_in
   836       val pos_pred = lev_back(*'*) pos
   837   	  val f_pred = Ctree.get_curr_formula (pt, pos_pred);
   838   	  val f_succ = Ctree.get_curr_formula (pt, pos);
   839       (*if*) f_succ = f_in (*else*);
   840   val NONE =
   841         (*case*) CAS_Cmd.input f_in (*of*);
   842 
   843 (*old* )     val (_, _, metID) = get_obj g_spec pt (par_pblobj pt p)
   844 (*old*)     val {scr = prog, ...} = MethodC.from_store metID
   845 (*old*)     val istate = get_istate_LI pt pos
   846 (*old*)     val ctxt = get_ctxt pt pos
   847   val LI.Found_Step (cstate'''''_', _(*istate*), _(*ctxt*)) = (*case*)
   848         LI.locate_input_term prog (pt, pos) istate ctxt f_in (*of*);
   849 "~~~~~ fun locate_input_term , args:"; val ((Rule.Prog _),  ((pt, pos) : Calc.T), (_ : Istate.T), (_ : Proof.context), tm)
   850   = (prog, (pt, pos), istate, ctxt, f_in);
   851 ( *old*)
   852 
   853 (*NEW*) LI.locate_input_term (pt, pos) f_in (*of*);
   854 "~~~~~ fun locate_input_term , args:"; val ((pt, pos), tm) = ((pt, pos), f_in);
   855    		val pos_pred = Pos.lev_back' pos (*f_pred ---"step pos cs"---> f_succ in appendFormula*)
   856 
   857   val ("ok", (_, _, cstate as (pt', pos'))) =
   858    		(*case*) compare_step ([], [], (pt, pos_pred)) tm (*of*);
   859 
   860 (*old* )
   861     Found_Step (cstate, get_istate_LI pt' pos', get_ctxt pt' pos')  (*return from locate_input_term*);
   862 ( *old*)
   863 (*NEW*)     Found_Step cstate (*return from locate_input_term*);
   864        (*LI.Found_Step ( *)cstate(*, _(*istate*), _(*ctxt*))( *return from locate_input_term*);
   865 "~~~~~ from fun locate_input_term\<longrightarrow>fun Step_Solve.by_term, return:"; val ("ok", (_(*use in DG !!!*), c, ptp as (_, p)))
   866   = (("ok" , ([], [], cstate (* already contains istate, ctxt *))));
   867 
   868     ("ok", ((*_ use in DG !!!,*) c, ptp(* as (_*), p))(*)*)(*return from Step_Solve.by_term*);
   869 "~~~~~ from fun Step_Solve.by_term\<longrightarrow>(fun appendFormula)!toplevel, return:"; val ("ok", (_(*use in DG !!!*), [], ptp''''' as (pt''''', p''''')))
   870   = ("ok", ([], [], ptp));
   871 
   872 (*fun me requires nxt...*)
   873     Step.do_next p''''' (ptp''''', []);
   874   val ("ok", ([(nxt'''''_' as Check_Postcond ["LINEAR", "univariate", "equation", "test"], _, _)], _,
   875     (pt'''''_', p'''''_'))) = Step.do_next p''''' (ptp''''', [])
   876 (*\\---------- appendFormula 1 "x = 1" \<longrightarrow> Step_Solve.inform \<longrightarrow> LI.locate_input_term ----------//*)
   877 
   878 (*//----- REPLACED BY appendFormula 1 "x = 1" \<longrightarrow> Step_Solve.inform \<longrightarrow> LI.locate_input_term -----\\* )
   879  (*[2], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Subproblem ("Test", ["LINEAR", "univariate", "equation", "test"])*)
   880  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Model_Problem*)
   881  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Add_Given "equality (- 1 + x = 0)"*)
   882  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Add_Given "solveFor x"*)
   883  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Add_Find "solutions x_i"*)
   884  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Specify_Theory "Test"*)
   885  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Specify_Problem ["LINEAR", "univariate", "equation", "test"]*)
   886  (*[3], Pbl*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Specify_Method ["Test", "solve_linear"]*)
   887  (*[3], Met*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Apply_Method ["Test", "solve_linear"]*)
   888  (*[3, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Rewrite_Set_Inst (["(''bdv'', x)"], "isolate_bdv")*)
   889  (*[3, 1], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Rewrite_Set "Test_simplify"*)
   890  (*[3, 2], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Check_Postcond ["LINEAR", "univariate", "equation", "test"]*)
   891 ( *\\----- REPLACED BY appendFormula 1 "x = 1" \<longrightarrow> Step_Solve.inform \<longrightarrow> LI.locate_input_term -----//*)
   892 
   893  (*[3], Res*)val (p,_,f,nxt,_,pt) = me nxt'''''_' p'''''_' [] pt'''''_'; (*nxt = Check_elementwise "Assumptions"*)
   894  (*[4], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Check_Postcond ["sqroot-test", "univariate", "equation", "test"]*)
   895  (*[], Res*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = End_Proof'*)
   896 
   897 (*/--- final test ---------------------------------------------------------------------------\\*)
   898 if p = ([], Res) andalso f2str f = "[x = 1]" andalso pr_ctree pr_short pt =
   899    ".    ----- pblobj -----\n" ^
   900    "1.   x + 1 = 2\n" ^
   901    "2.   x + 1 + - 1 * 2 = 0\n" ^
   902    "3.    ----- pblobj -----\n" ^
   903    "3.1.   - 1 + x = 0\n" ^
   904    "3.2.   x = 0 + - 1 * - 1\n" ^
   905    "3.2.1.   x = 0 + - 1 * - 1\n" ^
   906    "3.2.2.   x = 0 + 1\n" (*ATTENTION: see complete Calc below*)
   907 then case nxt of End_Proof' => () | _ => error "re-build: fun locate_input_term CHANGED 1"
   908 else error "re-build: fun locate_input_term CHANGED 2";
   909 
   910 Test_Tool.show_pt_tac pt; (*[
   911 ([], Frm), solve (x + 1 = 2, x)
   912 . . . . . . . . . . Apply_Method ["Test", "squ-equ-test-subpbl1"],
   913 ([1], Frm), x + 1 = 2
   914 . . . . . . . . . . Rewrite_Set "norm_equation",
   915 ([1], Res), x + 1 + - 1 * 2 = 0
   916 . . . . . . . . . . Rewrite_Set "Test_simplify",
   917 ([2], Res), - 1 + x = 0
   918 . . . . . . . . . . Subproblem (Test, ["LINEAR", "univariate", "equation", "test"]),
   919 ([3], Pbl), solve (- 1 + x = 0, x)
   920 . . . . . . . . . . Apply_Method ["Test", "solve_linear"],
   921 ([3,1], Frm), - 1 + x = 0
   922 . . . . . . . . . . Rewrite_Set_Inst ([(''bdv'', x)], "isolate_bdv"),
   923 ([3,1], Res), x = 0 + - 1 * - 1
   924 . . . . . . . . . . Derive Test_simplify,
   925 ([3,2,1], Frm), x = 0 + - 1 * - 1
   926 . . . . . . . . . . Rewrite ("#: - 1 * - 1 = 1", "- 1 * - 1 = 1"),
   927 ([3,2,1], Res), x = 0 + 1
   928 . . . . . . . . . . Rewrite ("radd_0", "0 + ?k = ?k"),
   929 ([3,2,2], Res), x = 1
   930 . . . . . . . . . . Tactic.input_to_string not impl. for ?!,
   931 ([3,2], Res), x = 1
   932 . . . . . . . . . . Check_Postcond ["LINEAR", "univariate", "equation", "test"],
   933 ([3], Res), [x = 1]
   934 . . . . . . . . . . Check_Postcond ["sqroot-test", "univariate", "equation", "test"],
   935 ([], Res), [x = 1]]*)
   936 
   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> ML \<open>
   946 \<close>
   947   ML_file "Interpret/step-solve.sml"
   948 
   949 ML_file "MathEngine/me-misc.sml"
   950   ML_file "MathEngine/fetch-tactics.sml"
   951 (** )ML_file "MathEngine/solve.sml"( *loops with eliminate ThmC.numerals_to_Free*)
   952   ML_file "MathEngine/step.sml"
   953 (** )
   954   ML_file "MathEngine/mathengine-stateless.sml"    (*!part. WN130804: +check Interpret/me.sml*)
   955                                    ( *loops with eliminate ThmC.numerals_to_Free*)
   956 ML_file "MathEngine/messages.sml"
   957   ML_file "MathEngine/states.sml"
   958 
   959   ML_file "BridgeLibisabelle/thy-present.sml"
   960   ML_file "BridgeLibisabelle/mathml.sml"           (*part.*)
   961   ML_file "BridgeLibisabelle/pbl-met-hierarchy.sml"
   962   ML_file "BridgeLibisabelle/thy-hierarchy.sml"
   963   ML_file "BridgeLibisabelle/interface-xml.sml"     (*TODO after 2009-2*)
   964 (** )ML_file "BridgeLibisabelle/interface.sml"( *loops with eliminate ThmC.numerals_to_Free*)
   965 (*WITHOUT inhibit exn WN1130621 Isabelle2012-->13 !thehier! THIS ERROR OCCURS:
   966   ... SAME ERROR HERE ON ISABELLE2012 AS IN ISAC ON ISABELLE2011*)
   967 
   968   ML_file "BridgeJEdit/parseC.sml"
   969   ML_file "BridgeJEdit/preliminary.sml"
   970 
   971   ML_file "Knowledge/delete.sml"
   972   ML_file "Knowledge/descript.sml"
   973   ML_file "Knowledge/simplify.sml"
   974   ML_file "Knowledge/poly.sml"
   975 ML \<open>
   976 \<close> ML \<open>
   977 \<close> ML \<open>
   978 \<close> ML \<open>
   979 \<close> ML \<open>
   980 \<close> ML \<open>
   981 \<close> ML \<open>
   982 \<close> ML \<open>
   983 \<close> ML \<open>
   984 \<close> ML \<open>
   985 \<close> ML \<open>
   986 \<close>
   987   ML_file "Knowledge/gcd_poly_ml.sml"
   988   ML_file "Knowledge/gcd_poly_winkler.sml" (*must be after gcd_poly_ml.sml: redefines functions*)
   989   ML_file "Knowledge/rational.sml"                                            (*Test_Isac_Short*)
   990   ML_file "Knowledge/equation.sml"
   991   ML_file "Knowledge/root.sml"
   992   ML_file "Knowledge/lineq.sml"
   993 
   994 (*ML_file "Knowledge/rooteq.sml"    some complicated equations not recovered from 2002 *)
   995 (*ML_file "Knowledge/rateq.sml"     some complicated equations not recovered----Test_Isac_Short*)
   996   ML_file "Knowledge/r/ootrat.sml"
   997   ML_file "Knowledge/rootrateq.sml"(*ome complicated equations not recovered from 2002 *)
   998 (*ML_file "Knowledge/partial_fractions.sml"  hangs with ML_system_64 = "true"---Test_Isac_Short*)
   999   ML_file "Knowledge/polyeq-1.sml"
  1000 (*ML_file "Knowledge/polyeq-2.sml"                                              Test_Isac_Short*)
  1001 (*ML_file "Knowledge/rlang.sml"     much to clean up, similar tests in other files     *)
  1002   ML_file "Knowledge/calculus.sml"
  1003   ML_file "Knowledge/trig.sml"
  1004 (*ML_file "Knowledge/logexp.sml"    not included as stuff for presentation of authoring*) 
  1005   ML_file "Knowledge/diff.sml"
  1006   ML_file "Knowledge/integrate.sml"
  1007   ML_file "Knowledge/eqsystem.sml"
  1008   ML_file "Knowledge/test.sml"
  1009   ML_file "Knowledge/polyminus.sml"
  1010   ML_file "Knowledge/vect.sml"
  1011   ML_file "Knowledge/diffapp.sml"        (* postponed to dev. specification | TP-prog. *)
  1012   ML_file "Knowledge/biegelinie-1.sml"
  1013 (*ML_file "Knowledge/biegelinie-2.sml"                                          Test_Isac_Short*)
  1014 (*ML_file "Knowledge/biegelinie-3.sml"                                          Test_Isac_Short*)
  1015 (** )ML_file "Knowledge/biegelinie-4.sml"( *loops with eliminate ThmC.numerals_to_Free*)
  1016 (** )ML_file "Knowledge/algein.sml"( *loops with eliminate ThmC.numerals_to_Free*)
  1017   ML_file "Knowledge/diophanteq.sml"
  1018 (*ML_file "Knowledge/inverse_z_transform.sml"hangs with ML_system_64 = "true"---Test_Isac_Short*)
  1019   ML_file "Knowledge/inssort.sml"
  1020   ML_file "Knowledge/isac.sml"
  1021   ML_file "Knowledge/build_thydata.sml"
  1022 
  1023   ML_file "Test_Code/test-code.sml"
  1024 
  1025 section \<open>further tests additional to src/.. files\<close>
  1026   ML_file "BridgeLibisabelle/use-cases.sml"
  1027 
  1028   ML \<open>"%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%";\<close>
  1029   ML \<open>"%%%%%%%%%%%%%%%%% end Test_Isac %%%%%%%%%%%%%%%%%%%%%%%%";\<close>
  1030   ML \<open>"%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%";\<close>
  1031 ML \<open>
  1032 \<close> ML \<open>
  1033 \<close> ML \<open>
  1034 \<close> ML \<open>
  1035 \<close> ML \<open>
  1036 \<close> ML \<open>
  1037 \<close> ML \<open>
  1038 \<close> ML \<open>
  1039 \<close> ML \<open>
  1040 \<close> ML \<open>
  1041 \<close> ML \<open>
  1042 \<close>
  1043 
  1044 section \<open>history of tests\<close>
  1045 text \<open>
  1046   Systematic regression tests have been introduced to isac development in 2003.
  1047   Sanity of the regression tests suffers from updates following Isabelle development,
  1048   which mostly exceeded the resources available in isac's development.
  1049 
  1050   The survey below shall support to efficiently use the tests for isac 
  1051   on different Isabelle versions. Conclusion in most cases will be: 
  1052 
  1053   !!! Use most recent tests or go back to the old notebook
  1054       with isac on Isabelle2002. !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
  1055 \<close>
  1056 
  1057 
  1058 subsection \<open>isac on Isabelle2017\<close>
  1059 subsubsection \<open>Summary of development\<close>
  1060 text \<open>
  1061   * Add further signatures, separate structures and cleanup respective files.
  1062   * Show feasibility of moving Isac programs to partial_function, i.e. to the function package.
  1063   * Clean theory dependencies.
  1064   * Start preparing shift from isac-java to Isabelle/jEdit.
  1065 \<close>
  1066 subsubsection \<open>State of tests: unchanged\<close>
  1067 subsubsection \<open>Changesets of begin and end\<close>
  1068 text \<open>
  1069   last changeset with Test_Isac 925fef0f4c81
  1070   first changeset with Test_Isac bbb414976dfe
  1071 \<close>
  1072 
  1073 subsection \<open>isac on Isabelle2015\<close>
  1074 subsubsection \<open>Summary of development\<close>
  1075 text \<open>
  1076   * Add signatures from top of thy-hierarchy down to Interpret (not ProgLang).
  1077     This complicates Test_Isac, see "Prepare running tests" above.
  1078   * Remove TTY interface.
  1079   * Re-activate insertion sort.
  1080 \<close>
  1081 subsubsection \<open>State of tests: unchanged\<close>
  1082 subsubsection \<open>Changesets of begin and end\<close>
  1083 text \<open>
  1084   last changeset with Test_Isac 2f1b2854927a
  1085   first changeset with Test_Isac ???
  1086 \<close>
  1087 
  1088 subsection \<open>isac on Isabelle2014\<close>
  1089 subsubsection \<open>Summary of development\<close>
  1090 text \<open>
  1091   migration from "isabelle tty" --> libisabelle
  1092 \<close>
  1093 
  1094 subsection \<open>isac on Isabelle2013-2\<close>
  1095 subsubsection \<open>Summary of development\<close>
  1096 text \<open>
  1097   reactivated context_thy
  1098 \<close>
  1099 subsubsection \<open>State of tests\<close>
  1100 text \<open>
  1101   TODO
  1102 \<close>
  1103 subsubsection \<open>Changesets of begin and end\<close>
  1104 text \<open>
  1105   TODO
  1106   :
  1107   : isac on Isablle2013-2
  1108   :
  1109   Changeset: 55318 (03826ceb24da) merged
  1110   User: Walther Neuper <neuper@ist.tugraz.at>
  1111   Date: 2013-12-12 14:27:37 +0100 (7 minutes)
  1112 \<close>
  1113 
  1114 subsection \<open>isac on Isabelle2013-1\<close>
  1115 subsubsection \<open>Summary of development\<close>
  1116 text \<open>
  1117   Isabelle2013-1 was replaced within a few weeks due to problems with the document model;
  1118   no significant development steps for ISAC.
  1119 \<close>
  1120 subsubsection \<open>State of tests\<close>
  1121 text \<open>
  1122   See points in subsection "isac on Isabelle2011", "State of tests".
  1123 \<close>
  1124 subsubsection \<open>Changesets of begin and end\<close>
  1125 text \<open>
  1126   Changeset: 55283 (d6e9a34e7142) notes for resuming work on Polynomial.thy
  1127   User: Walther Neuper <neuper@ist.tugraz.at>
  1128   Date: 2013-12-03 18:13:31 +0100 (8 days)
  1129   :
  1130   : isac on Isablle2013-1
  1131   :
  1132   Changeset: 55279 (130688f277ba) Isabelle2013 --> 2013-1: Test_Isac perfect
  1133   User: Walther Neuper <neuper@ist.tugraz.at>
  1134   Date: 2013-11-21 18:12:17 +0100 (2 weeks)
  1135 
  1136 \<close>
  1137 
  1138 subsection \<open>isac on Isabelle2013\<close>
  1139 subsubsection \<open>Summary of development\<close>
  1140 text \<open>
  1141   # Oct.13: replaced "axioms" by "axiomatization"
  1142   # Oct.13: Mathias Lehnfeld started removing Unsynchornized.ref
  1143   # Sep.13: integrated gcd_poly (functional, without Unsychronized.ref) into
  1144     simplification of multivariate rationals (without improving the rulesets involved).
  1145 \<close>
  1146 subsubsection \<open>Run tests\<close>
  1147 text \<open>
  1148   Is standard now; this subsection will be discontinued under Isabelle2013-1
  1149 \<close>
  1150 subsubsection \<open>State of tests\<close>
  1151 text \<open>
  1152   See points in subsection "isac on Isabelle2011", "State of tests".
  1153   # re-activated listC.sml
  1154 \<close>
  1155 subsubsection \<open>Changesets of begin and end\<close>
  1156 text \<open>
  1157   changeset 52174:8b055b17bd84 --- removed all code concerned with "castab = Unsynchronized.ref"
  1158   User: Mathias Lehnfeld <s1210629013@students.fh-hagenberg.at>
  1159   Date: Tue Nov 19 22:23:30 2013 +0000
  1160   :
  1161   : isac on Isablle2013 
  1162   :
  1163   Changeset: 52061 (4ecea2fcdc2c) --- Build_Isac.thy runs on Isabelle2013
  1164   User: Walther Neuper <neuper@ist.tugraz.at>
  1165   Date: 2013-07-15 08:28:50 +0200 (4 weeks)
  1166 \<close>
  1167 
  1168 subsection \<open>isac on Isabelle2012\<close>
  1169 subsubsection \<open>Summary of development\<close>
  1170 text \<open>
  1171   isac on Isabelle2012 is considered just a transitional stage
  1172   within the update from Isabelle2011 to Isabelle2013; thus no further development of isac;
  1173   For considerations on the transition see 
  1174   $ISABELLE_ISAC/Knowledge/Build_Thydata/thy, section "updating isac..".
  1175 \<close>
  1176 subsubsection \<open>Run tests\<close>
  1177 text \<open>
  1178 $ cd /usr/local/isabisac12/
  1179 $ ./bin/isabelle jedit -l HOL src/Tools/isac/Build_Isac.thy
  1180 $ ./bin/isabelle jedit -l Isac test/Tools/isac/Test_Isac.thy
  1181 \<close>
  1182 subsubsection \<open>State of tests\<close>
  1183 text \<open>
  1184   At least the tests from isac on Isabelle2011 run again.
  1185   However, Test_Isac.thy shows erratic behaviour; no errors are obtained when scrolling 
  1186   in parallel with evaluation.
  1187 
  1188   Counting "error in kernel" for Frontend/interface.sml (the tests considered most significant)
  1189   yields 69 hits, some of which were already present before Isabelle2002-->2009-2
  1190   (i.e. on the old notebook from 2002).
  1191 
  1192   Now many tests with (*...=== inhibit exn ...*) give a reason or at least the origin:
  1193   # === inhibit exn WN1130621 Isabelle2012-->13 !thehier! === ...see Build_Thydata.thy
  1194   # === inhibit exn AK110726 === ...reliable work by Alexander Kargl, most likely go back to 2002
  1195   # === inhibit exn WN1130701 broken at Isabelle2002 --> 2009-2 === , most likely go back to 2002
  1196   Reasons for outcommented tests are also found in Test_Isac.thy near the respective file.sml.
  1197 
  1198   Some tests have been re-activated (e.g. error patterns, fill patterns).
  1199 \<close>
  1200 subsubsection \<open>Changesets of begin and end\<close>
  1201 text \<open>
  1202   Changeset: 52051 (35751d90365e) end of improving tests for isac on Isabelle2012
  1203   User: Walther Neuper <neuper@ist.tugraz.at>
  1204   Date: 2013-07-11 16:58:31 +0200 (4 weeks)
  1205   :
  1206   : isac on Isablle2012 
  1207   :
  1208   Changeset: 48757 (74eb3dfc33cc) updated src from Isabelle2011 to Isabelle2012
  1209   User: Walther Neuper <neuper@ist.tugraz.at>
  1210   Date: 2012-09-24 18:35:13 +0200 (8 months)
  1211   ------------------------------------------------------------------------------
  1212   Changeset: 48756 (7443906996a8) merged
  1213   User: Walther Neuper <neuper@ist.tugraz.at>
  1214   Date: 2012-09-24 18:15:49 +0200 (8 months)
  1215 \<close>
  1216 
  1217 subsection \<open>isac on Isabelle2011\<close>
  1218 subsubsection \<open>Summary of development\<close>
  1219 text \<open>
  1220   isac's mathematics engine has been extended by two developments:
  1221   (1) Isabelle's contexts were introduced by Mathias Lehnfeld
  1222   (2) Z_Transform was introduced by Jan Rocnik, which revealed
  1223     further errors introduced by (1).
  1224   (3) "error patterns" were introduced by Gabriella Daroczy
  1225   Regressions tests have been added for all of these.
  1226 \<close>
  1227 subsubsection \<open>Run tests\<close>
  1228 text \<open>
  1229   $ cd /usr/local/isabisac11/
  1230   $ ./bin/isabelle jedit -l HOL src/Tools/isac/Build_Isac.thy
  1231   $ ./bin/isabelle jedit -l Isac test/Tools/isac/Test_Isac.thy
  1232 \<close>
  1233 subsubsection \<open>State of tests\<close>
  1234 text \<open>
  1235   Systematic efforts outcommented less significant tests by (*...=== inhibit exn ...*) 
  1236   and sometimes give reasons for failing tests.
  1237   (*...=== inhibit exn AK...*) was done by Alexander Kargl; this is reliable
  1238   work, some of which couldn't be revised (and renamed) by WN and thus survived some time.
  1239 
  1240   The most signification tests (in particular Frontend/interface.sml) run,
  1241   however, many "error in kernel" are not caught by an exception.
  1242   ------------------------------------------------------------------------------
  1243   After the changeset below Test_Isac worked with check_unsynchronized_ref ():
  1244   ------------------------------------------------------------------------------
  1245   Changeset: 42457 (ca691a84b81a) PROVISIONALLY MADE TESTS RUN with Unsynchronized.ref
  1246   User: Walther Neuper <neuper@ist.tugraz.at>
  1247   Date: 2012-08-06 10:38:11 +0200 (11 months)
  1248 
  1249 
  1250   The list below records TODOs while producing an ISAC kernel for 
  1251   gdaroczy and jrocnik, wich could NOT be done before all tests are RUNNING
  1252   (so to be resumed with Isabelle2013-1):
  1253   ############## WNxxxxxx.TODO can be found in sources ##############
  1254   --------------------------------------------------------------------------------
  1255   WN111013.TODO: lots of cleanup/removal in test/../Test.thy
  1256   --------------------------------------------------------------------------------
  1257   WN111013.TODO: remove concept around "fun implicit_take", lots of troubles with 
  1258   this special case (see) --- why not nxt = Model_Problem here ? ---
  1259   --------------------------------------------------------------------------------
  1260   WN111014.TODO calculate_Poly < calculate_Rational < calculate_RootRat, see test/
  1261   ... FIRST redesign 
  1262   # simplify_* , *_simp_* 
  1263   # norm_* 
  1264   # calc_* , calculate_*  ... require iteration over all rls ...
  1265   ... see --- val rls = calculate_RootRat > calculate_Rational --- CONTINUE !
  1266   --------------------------------------------------------------------------------
  1267   WN111014.TODO fun prep_rls | !!!use this function in ruleset' := !!!
  1268   --------------------------------------------------------------------------------
  1269   WN120314 changeset a393bb9f5e9f drops root equations.
  1270   see test/Tools/isac/Knowledge/rootrateq.sml 
  1271   --------------------------------------------------------------------------------
  1272   WN120317.TODO changeset 977788dfed26 dropped rateq:
  1273   # test --- repair NO asms from rls RatEq_eliminate --- shows error from 2002
  1274   # test --- solve (1/x = 5, x) by me --- and --- x / (x  \<up>  2 - 6 * x + 9) - ...:    
  1275     investigation Check_elementwise stopped due to too much effort finding out,
  1276     why Check_elementwise worked in 2002 in spite of the error.
  1277   --------------------------------------------------------------------------------
  1278   WN120317.TODO postponed test/../ratinal,ratinal2.sml to joint work with dmeindl 
  1279   --------------------------------------------------------------------------------
  1280   WN120317.TODO found by test --- interSteps for Schalk 299a --- that 
  1281     NO test with 'interSteps' is checked properly (with exn on changed behaviour)
  1282   --------------------------------------------------------------------------------
  1283   WN120317.TODO test --- Matthias Goldgruber 2003 rewrite orders --- has
  1284     a newly outcommented test where rewrite_set_ make_polynomial --> NONE
  1285   --------------------------------------------------------------------------------
  1286   WN120320.TODO check-improve rlsthmsNOTisac:
  1287   DONE make test --- old compute rlsthmsNOTisac by eq_thmI'
  1288   DONE compare rlsthmsNOTisac in thms-survey-Isa02-Isa09-2.sml .. Isac.thy 
  1289   FOUND 120321: Theory.axioms_of doesnt find LENGTH_CONS etc, thus are in Isab
  1290   # mark twice thms (in isac + (later) in Isabelle) in Isac.thy
  1291   --------------------------------------------------------------------------------
  1292   WN120320.TODO rlsthmsNOTisac: replace twice thms  \<up> 
  1293   --------------------------------------------------------------------------------
  1294   WN120320.TODO rlsthmsNOTisac: reconsider design of sym_* thms, see test
  1295   --- OLD compute rlsthmsNOTisac by eq_thmID ---: some are in isab, some in isac.
  1296   --------------------------------------------------------------------------------
  1297   WN120321.TODO rearrange theories:
  1298     Knowledge
  1299       :
  1300       Prog_Expr.thy
  1301       ///Input_Descript.thy --> ProgLang
  1302       Delete.thy   <--- first_Knowledge_thy (*mv to Prog_Expr.thy*)
  1303     ProgLang: restructure Build_Isac.thy such that no xmlsrc.thy, ProgLang.thy
  1304           Interpret.thy are generated (simplifies xml structure for theories)
  1305       Program.thy
  1306       Tools.thy
  1307       ListC.thy    <--- first_Proglang_thy
  1308   --------------------------------------------------------------------------------
  1309   WN120321.TODO reanimate test/../simplify.sml: CAS-command Simplify
  1310       EXN "simplify.sml: diff.behav. CAScmd: Simplify (2*a + 3*a)"
  1311   broken during work on thy-hierarchy
  1312   --------------------------------------------------------------------------------
  1313   WN120321.TODO LAST in IsacScripts + in IsacKnowledge, see
  1314   test --- the_hier (get_thes ()) (collect_thydata ())---
  1315   --------------------------------------------------------------------------------
  1316   WN120405a.TODO src/../pbl-met-hierarchy.sml: fun pbl2term--> Isac' instead Isac_
  1317   !!add mutual crossreferences to ?fun headline??? where the same has to be done:
  1318   !!OR BETTER: use only 2 functions for adding/removing "'" to/from thy!!
  1319   --------------------------------------------------------------------------------
  1320   WN120411 scanning html representation of newly generated knowledge:
  1321   * thy:
  1322   ** Theorems: only "Proof of the theorem" (correct!)
  1323                and "(c) isac-team (math-autor)"
  1324   ** Rulesets: only "Identifier:///"
  1325                and "(c) isac-team (math-autor)"
  1326   ** IsacKnowledge: link to dependency graph (which needs to be created first)
  1327   ** IsacScripts --> ProgramLanguage
  1328   *** Tools: Theorems: arity_type_cp, arity_type_nam, ... wegnehmen
  1329   
  1330   * pbl: OK !?!
  1331   * met: OK !?!
  1332   * exp: 
  1333   ** Z-Transform is missing !!!
  1334   ** type-constraints !!!
  1335   --------------------------------------------------------------------------------
  1336   WN120417: merging xmldata revealed:
  1337   ..............NEWLY generated:........................................
  1338   <THEOREMDATA>
  1339     <GUH> thy_isab_Fun-thm-o_apply </GUH>
  1340     <STRINGLIST>
  1341       <STRING> Isabelle </STRING>
  1342       <STRING> Fun </STRING>
  1343       <STRING> Theorems </STRING>
  1344       <STRING> o_apply </STRING>
  1345     </STRINGLIST>
  1346       <MATHML>
  1347         <ISA> (?f o ?g) ?x = ?f (?g ?x) </ISA>
  1348       </MATHML>  <PROOF>
  1349       <EXTREF>
  1350         <TEXT> Proof of the theorem </TEXT>
  1351         <URL> http://www.ist.tugraz.at/projects/isac/www/kbase/thy/browser_info/HOL/HOL-Real/Isac/Fun.html </URL>
  1352       </EXTREF>
  1353     </PROOF>
  1354     <EXPLANATIONS> </EXPLANATIONS>
  1355     <MATHAUTHORS>
  1356       <STRING> Isabelle team, TU Munich </STRING>
  1357     </MATHAUTHORS>
  1358     <COURSEDESIGNS>
  1359     </COURSEDESIGNS>
  1360   </THEOREMDATA>
  1361   ..............OLD FORMAT:.............................................
  1362   <THEOREMDATA>
  1363     <GUH> thy_isab_Fun-thm-o_apply </GUH>
  1364     <STRINGLIST>
  1365       <STRING> Isabelle </STRING>
  1366       <STRING> Fun </STRING>
  1367       <STRING> Theorems </STRING>
  1368       <STRING> o_apply </STRING>
  1369     </STRINGLIST>
  1370     <THEOREM>
  1371       <ID> o_apply </ID>
  1372       <MATHML>
  1373         <ISA> (?f o ?g) ?x = ?f (?g ?x) </ISA>
  1374       </MATHML>
  1375     </THEOREM>
  1376     <PROOF>
  1377       <EXTREF>
  1378         <TEXT> Proof of the theorem </TEXT>
  1379         <URL> http://www.ist.tugraz.at/projects/isac/www/kbase/thy/browser_info/HOL/HOL-Real/Isac/Fun.html </URL>
  1380       </EXTREF>
  1381     </PROOF>
  1382     <EXPLANATIONS> </EXPLANATIONS>
  1383     <MATHAUTHORS>
  1384       <STRING> Isabelle team, TU Munich </STRING>
  1385     </MATHAUTHORS>
  1386     <COURSEDESIGNS>
  1387     </COURSEDESIGNS>
  1388   </THEOREMDATA>
  1389   --------------------------------------------------------------------------------
  1390 \<close>
  1391 subsubsection \<open>Changesets of begin and end\<close>
  1392 text \<open>
  1393   isac development was done between these changesets:
  1394   ------------------------------------------------------------------------------
  1395   Changeset: 42519 (1f3b4270363e) meeting dmeindl: added missing files
  1396   User: Walther Neuper <neuper@ist.tugraz.at>
  1397   Date: 2012-09-24 16:39:30 +0200 (8 months)
  1398   :
  1399   : isac on Isablle2011
  1400   :
  1401   Changeset:41897 (355be7f60389) merged isabisac with Isabelle2011
  1402   Branch: decompose-isar 
  1403   User: Walther Neuper <neuper@ist.tugraz.at>
  1404   Date: 2011-02-25 13:04:56 +0100 (2011-02-25)
  1405   ------------------------------------------------------------------------------
  1406 \<close>
  1407 
  1408 subsection \<open>isac on Isabelle2009-2\<close>
  1409 subsubsection \<open>Summary of development\<close>
  1410 text \<open>
  1411   In 2009 the update of isac from Isabelle2002 started with switching from CVS to hg.
  1412   The update was painful (bridging 7 years of Isabelle development) and cut short 
  1413   due to the start of introducing Isabelle's contexts (Mathias Lehnfeld) and
  1414   going on to Isabelle2011 although most of the tests did not run.
  1415 \<close>
  1416 subsubsection \<open>Run tests\<close>
  1417 text \<open>
  1418   WN131021 this is broken by installation of Isabelle2011/12/13,
  1419   because all these write their binaries to ~/.isabelle/heaps/..
  1420 
  1421   $ cd /usr/local/isabisac09-2/
  1422   $ ./bin/isabelle emacs -l HOL src/Tools/isac/Build_Isac.thy
  1423   $ ./bin/isabelle emacs -l Isac src/Tools/isac/Test_Isac.thy
  1424   NOT THE RIGHT VERSION.....    test/Tools/isac/Test_Isac.thy !!!
  1425 \<close>
  1426 subsubsection \<open>State of tests\<close>
  1427 text \<open>
  1428   Most tests are broken by the update from Isabelle2002 to Isabelle2009-2.
  1429 \<close>
  1430 subsubsection \<open>Changesets of begin and end\<close>
  1431 text \<open>
  1432   isac development was done between these changesets:
  1433   ------------------------------------------------------------------------------
  1434   Changeset: 38115 (940a5feea094) Little improvements of isac-plugin
  1435   Branch: decompose-isar 
  1436   User: Marco Steger <m.steger@student.tugraz.at>
  1437   Date: 2011-02-06 18:30:28 +0100 (2011-02-06)
  1438   :
  1439   : isac on Isablle2009-2
  1440   :
  1441   Changeset: 37870 (5100a9c3abf8) created branch isac-from-Isabelle2009-2
  1442   Branch: isac-from-Isabelle2009-2 
  1443   User: Walther Neuper <neuper@ist.tugraz.at>
  1444   Date: 2010-07-21 09:59:35 +0200 (2010-07-21)
  1445   ------------------------------------------------------------------------------
  1446 \<close>
  1447 
  1448 subsection \<open>isac on Isabelle2002\<close>
  1449 subsubsection \<open>Summary of development\<close>
  1450 text \<open>
  1451   From 1999 to 2010 all the basic functionality (except "ctxt" and "error pattern")
  1452   of isac's mathematics engine has been implemented.
  1453 \<close>
  1454 subsubsection \<open>Run tests\<close>
  1455 subsubsection \<open>State of tests\<close>
  1456 text \<open>
  1457   All tests work on an old notebook (the right PolyML coudn't be upgraded to more
  1458   recent Linux versions)
  1459 \<close>
  1460 subsubsection \<open>Changesets of begin and end\<close>
  1461 text \<open>
  1462   Up to the above Mercurial changeset 5100a9c3abf8 isac used CVS;
  1463   see https://intra.ist.tugraz.at/hg/isac containing a conversion to Mercurial.
  1464 \<close>
  1465 
  1466 end
  1467 (*========== inhibit exn 130719 Isabelle2013 ===================================
  1468 ============ inhibit exn 130719 Isabelle2013 =================================*)
  1469 
  1470 (*-.-.-.-.-.-isolate response.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.
  1471   -.-.-.-.-.-.-isolate response.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.*)
  1472