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