test/Tools/isac/Test_Isac.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Wed, 05 Oct 2016 13:09:54 +0200
changeset 59248 5eba5e6d5266
parent 59232 a757e8f2fe6c
child 59252 7d3dbc1171ff
permissions -rwxr-xr-x
cleaned tests from autoCalculate' (removed ')

Notes
(1) autoCalculate' made Test_Some.thy complicated (had to be redefined)
(2) "fun autoCalculate'" was scattered over several tests -- probably due to (1)
(3) autoCalculate' has been put out of service in 5127be395ea1
(4) M.Lehnfeld finished his work with 8e3f73e1e3a3:
fun autoCalculate (cI:calcID) auto = Future.fork ... in isac/../interface.sml
     1 (* Title:  All tests on isac (some outcommented since Isabelle2002-->2009-2)
     2    Author: Walther Neuper 101001
     3    (c) copyright due to lincense terms.
     4 
     5    Isac's tests are organised parallel to sources: 
     6      "~~/test/Tools/isac" has same directory structure as "~~/src/Tools/isac"
     7    plus
     8      ~~/test/Tools/isac/ADDTESTS
     9      ~~/test/Tools/isac/Minisubpbl: the Lucas-Interpreter's core functionality
    10 
    11 $ cd /usr/local/isabisac/
    12 $ ./bin/isabelle jedit -l Isac test/Tools/isac/Test_Isac.thy &
    13 *)
    14 
    15 (* !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!! *)
    16 (* !!!!! say "OK" to the popup asking for theories to be loaded and !!!!!!!!!!!!! *)
    17 (* !!!!! watch the <Theories> window for errors in the "imports" below !!!!!!!!!! *)
    18 (* !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!! *)
    19 
    20 theory Test_Isac imports Build_Thydata
    21   "ADDTESTS/accumulate-val/Thy_All"
    22   "ADDTESTS/Ctxt"
    23   "ADDTESTS/test-depend/Build_Test"
    24   "ADDTESTS/All_Ctxt"
    25   "ADDTESTS/Test_Units"
    26   "ADDTESTS/course/phst11/T1_Basics"
    27   "ADDTESTS/course/phst11/T2_Rewriting"
    28   "ADDTESTS/course/phst11/T3_MathEngine"
    29   "ADDTESTS/file-depend/BuildC_Test"
    30   "ADDTESTS/session-get_theory/Foo"
    31 (*"ADDTESTS/course/SignalProcess/Build_Inverse_Z_Transform"
    32    ADDTESTS/------------------------------------------- see end of tests *)
    33 (*"~~/test/Pure/Isar/Test_Parsers"           dropped Isabelle2014-->2015 *)
    34 (*"~~/test/Pure/Isar/Pure/Isar/Struct_Deriv" lost at update 2009-2-->2011*)
    35   "~~/test/Pure/Isar/Test_Parse_Term"
    36   "~~/test/HOL/Library/Test_Polynomial"
    37   "~~/test/Tools/isac/Interpret/ptyps"       (* setup for ptyps.sml    *)
    38   "~~/test/Tools/isac/ProgLang/calculate"    (* setup for calculate.sml*)
    39   "~~/test/Tools/isac/ProgLang/scrtools"     (* setup for scrtools.sml *)
    40   "~~/test/Tools/isac/Knowledge/integrate"   (* setup for integrate.sml*)
    41 
    42   "~~/src/Tools/isac/Knowledge/GCD_Poly_OLD" (*not imported by Isac.thy*)
    43   "~~/src/Tools/isac/Knowledge/GCD_Poly_FP"  (*not imported by Isac.thy*)
    44 
    45 begin
    46 
    47 ML {*
    48   KEStore_Elems.set_ref_thy @{theory};
    49   (*fun autoCalculate' cI auto = autoCalculate cI auto (*|> Future.join*)*);
    50 *}
    51 
    52 section {* trials with Isabelle's functions *}
    53   ML {*"%%%%%%%%%%%%%%%%% start Isabelle %%%%%%%%%%%%%%%%%%%%%%%";*}
    54   ML_file "~~/test/Pure/General/alist.ML"
    55   ML_file "~~/test/Pure/General/basics.ML"
    56   ML_file "~~/test/Pure/General/scan.ML"
    57   ML_file "~~/test/Pure/PIDE/xml.ML"
    58   ML {*"%%%%%%%%%%%%%%%%% end Isabelle %%%%%%%%%%%%%%%%%%%%%%%%%";*}
    59 
    60 section {* test ML Code of isac *}
    61   ML {*"%%%%%%%%%%%%%%%%% start ProgLang.thy %%%%%%%%%%%%%%%%%%%";*}
    62   ML_file          "library.sml"
    63   ML_file          "calcelems.sml"
    64   ML_file          "kestore.sml"
    65   ML_file "ProgLang/termC.sml"
    66   ML_file "ProgLang/calculate.sml"      (* requires setup from calculate.thy *)
    67   ML_file "ProgLang/rewrite.sml"
    68   ML_file "ProgLang/listC.sml"
    69   ML_file "ProgLang/scrtools.sml"
    70   ML_file "ProgLang/tools.sml"
    71   ML {*"%%%%%%%%%%%%%%%%% end ProgLang.thy %%%%%%%%%%%%%%%%%%%%%";*}
    72   ML {*"%%%%%%%%%%%%%%%%% start Minisubpbl %%%%%%%%%%%%%%%%%%%%%";*}
    73   ML_file "Minisubpbl/000-comments.sml"
    74   ML_file "Minisubpbl/100-init-rootpbl.sml"
    75   ML_file "Minisubpbl/150-add-given.sml"
    76   ML_file "Minisubpbl/200-start-method.sml"
    77   ML_file "Minisubpbl/300-init-subpbl.sml"
    78   ML_file "Minisubpbl/400-start-meth-subpbl.sml"
    79   ML_file "Minisubpbl/490-nxt-Check_Postcond.sml"
    80   ML_file "Minisubpbl/500-met-sub-to-root.sml"
    81   ML_file "Minisubpbl/530-error-Check_Elementwise.sml"
    82   ML_file "Minisubpbl/600-postcond.sml"
    83   ML {*"%%%%%%%%%%%%%%%%% end Minisubpbl.thy %%%%%%%%%%%%%%%%%%%%%%%";*}
    84   ML {*"%%%%%%%%%%%%%%%%% start Interpret.thy %%%%%%%%%%%%%%%%%%";*}
    85   ML_file "Interpret/mstools.sml"
    86   ML_file "Interpret/ctree.sml"         (*!...!see(25)*)
    87   ML_file "Interpret/ptyps.sml"         (* requires setup from ptyps.thy *)
    88   ML {* (*check_unsynchronized_ref (); ==== trick on error: CUT AND PASTE THIS LINE =========*) *}
    89 (*TRICK DOESN'T WORK: get_pbt not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*)
    90   ML_file "Interpret/generate.sml"
    91 (*WITHOUT inhibit exn WN1130621 Isabelle2012-->13 !thehier! THIS ERROR OCCURS:
    92   ... SAME ERROR HERE ON ISABELLE2012 AS IN ISAC ON ISABELLE2011*)
    93   ML_file "Interpret/calchead.sml"
    94   ML_file "Interpret/appl.sml"          (*complete "WEGEN INTERMED TESTCODE"                          *)
    95   ML_file "Interpret/rewtools.sml"
    96   ML_file "Interpret/script.sml"
    97   ML_file "Interpret/solve.sml"
    98   ML_file "Interpret/inform.sml"
    99 (*WITHOUT inhibit exn WN1130621 Isabelle2012-->13 !thehier! THIS ERROR OCCURS:
   100   ... SAME ERROR HERE ON ISABELLE2012 AS IN ISAC ON ISABELLE2011*)
   101   ML_file "Interpret/mathengine.sml"    (*!part. WN130804: +check Interpret/me.sml*)
   102   ML {*"%%%%%%%%%%%%%%%%% end Interpret.thy %%%%%%%%%%%%%%%%%%%%";*}
   103   ML {*"%%%%%%%%%%%%%%%%% start xmlsrc.thy %%%%%%%%%%%%%%%%%%%%%";*}
   104   ML_file "xmlsrc/mathml.sml"           (*part.*)
   105   ML_file "xmlsrc/datatypes.sml"        (*TODO/part.*)
   106   ML_file "xmlsrc/pbl-met-hierarchy.sml"(*TODO after 2009-2/part.*)
   107   ML_file "xmlsrc/thy-hierarchy.sml"
   108   ML_file "xmlsrc/interface-xml.sml"     (*TODO after 2009-2*)
   109   ML {*"%%%%%%%%%%%%%%%%% end xmlsrc.thy %%%%%%%%%%%%%%%%%%%%%%%";*}
   110   ML {*"%%%%%%%%%%%%%%%%% start Frontend.thy %%%%%%%%%%%%%%%%%%%";*}
   111   ML_file "Frontend/messages.sml"
   112   ML_file "Frontend/states.sml"
   113   ML_file "Frontend/interface.sml"
   114   ML_file "Frontend/use-cases.sml"
   115 (*WITHOUT inhibit exn WN1130621 Isabelle2012-->13 !thehier! THIS ERROR OCCURS:
   116   ... SAME ERROR HERE ON ISABELLE2012 AS IN ISAC ON ISABELLE2011*)
   117   ML_file          "print_exn_G.sml"
   118   ML {*"%%%%%%%%%%%%%%%%% end Frontend.thy %%%%%%%%%%%%%%%%%%%%%";*}
   119   ML {*"%%%%%%%%%%%%%%%%% start Knowledge %%%%%%%%%%%%%%%%%%%%%%";*}
   120   ML_file "Knowledge/delete.sml"
   121   ML_file "Knowledge/descript.sml"
   122   ML_file "Knowledge/atools.sml"
   123   ML_file "Knowledge/simplify.sml"
   124   ML_file "Knowledge/poly.sml"
   125   ML_file "Knowledge/gcd_poly_ml.sml"
   126   ML_file "Knowledge/gcd_poly_winkler.sml" (*must be after gcd_poly_ml.sml: redefines functions*)
   127   ML_file "Knowledge/rational.sml"
   128   ML_file "Knowledge/equation.sml"
   129   ML_file "Knowledge/root.sml"
   130   ML_file "Knowledge/lineq.sml"
   131 (*ML_file "Knowledge/rooteq.sml"    some complicated equations not recovered from 2002 *)
   132   ML_file "Knowledge/rateq.sml"   (*some complicated equations not recovered from 2002 *)
   133   ML_file "Knowledge/rootrat.sml"
   134   ML_file "Knowledge/rootrateq.sml"(*ome complicated equations not recovered from 2002 *)
   135   ML_file "Knowledge/partial_fractions.sml"
   136   ML_file "Knowledge/polyeq.sml"
   137 (*ML_file "Knowledge/rlang.sml"     much to clean up, similar tests in other files     *)
   138   ML_file "Knowledge/calculus.sml"
   139   ML_file "Knowledge/trig.sml"
   140 (*ML_file "Knowledge/logexp.sml"    not included as stuff for presentation of authoring*) 
   141   ML_file "Knowledge/diff.sml"
   142   ML_file "Knowledge/integrate.sml"
   143   ML_file "Knowledge/eqsystem.sml"
   144   ML_file "Knowledge/test.sml"
   145   ML_file "Knowledge/polyminus.sml"
   146   ML_file "Knowledge/vect.sml"
   147   ML_file "Knowledge/diffapp.sml"        (* postponed to dev. specification | TP-prog. *)
   148   ML_file "Knowledge/biegelinie.sml"
   149   ML_file "Knowledge/algein.sml"
   150   ML_file "Knowledge/diophanteq.sml"
   151   ML_file "Knowledge/Inverse_Z_Transform/inverse_z_transform.sml"
   152   ML_file "Knowledge/inssort.sml"
   153   ML_file "Knowledge/isac.sml"
   154   ML_file "Knowledge/build_thydata.sml"
   155   ML {*"%%%%%%%%%%%%%%%%% end Knowledge %%%%%%%%%%%%%%%%%%%%%%%%";*}
   156   ML {*"%%%%%%%%%%%%%%%%% start ADDTESTS %%%%%%%%%%%%%%%%%%%%%%%";*}
   157   ML_file "~~/test/Tools/isac/ADDTESTS/libisabelle/mini-test.sml"
   158   ML_file "~~/test/Tools/isac/ADDTESTS/libisabelle/protocol.sml"
   159   ML {*"%%%%%%%%%%%%%%%%% end ADDTESTS %%%%%%%%%%%%%%%%%%%%%%%%%";*}
   160 
   161   ML {*"%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%";*}
   162   ML {*"%%%%%%%%%%%%%%%%% end Test_Isac %%%%%%%%%%%%%%%%%%%%%%%%";*}
   163   ML {*"%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%";*}
   164 
   165 section {* history of tests *}
   166 text {*
   167   Systematic regression tests have been introduced to isac development in 2003.
   168   Sanity of the regression tests suffers from updates following Isabelle development,
   169   which mostly exceeded the resources available in isac's development.
   170 
   171   The survey below shall support to efficiently use the tests for isac 
   172   on different Isabelle versions. Conclusion in most cases will be: 
   173 
   174   !!! Use most recent tests or go back to the old notebook
   175       with isac on Isabelle2002. !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
   176 *}
   177 
   178 subsection {* isac on Isabelle2014 *}
   179 subsubsection {* Summary of development *}
   180 text {*
   181   migration from "isabelle tty" --> libisabelle
   182 *}
   183 
   184 subsection {* isac on Isabelle2013-2 *}
   185 subsubsection {* Summary of development *}
   186 text {*
   187   reactivated context_thy
   188 *}
   189 subsubsection {* State of tests *}
   190 text {*
   191   TODO
   192 *}
   193 subsubsection {* Changesets of begin and end *}
   194 text {*
   195   TODO
   196   :
   197   : isac on Isablle2013-2
   198   :
   199   Changeset: 55318 (03826ceb24da) merged
   200   User: Walther Neuper <neuper@ist.tugraz.at>
   201   Date: 2013-12-12 14:27:37 +0100 (7 minutes)
   202 *}
   203 
   204 subsection {* isac on Isabelle2013-1 *}
   205 subsubsection {* Summary of development *}
   206 text {*
   207   Isabelle2013-1 was replaced within a few weeks due to problems with the document model;
   208   no significant development steps for ISAC.
   209 *}
   210 subsubsection {* State of tests *}
   211 text {*
   212   See points in subsection "isac on Isabelle2011", "State of tests".
   213 *}
   214 subsubsection {* Changesets of begin and end *}
   215 text {*
   216   Changeset: 55283 (d6e9a34e7142) notes for resuming work on Polynomial.thy
   217   User: Walther Neuper <neuper@ist.tugraz.at>
   218   Date: 2013-12-03 18:13:31 +0100 (8 days)
   219   :
   220   : isac on Isablle2013-1
   221   :
   222   Changeset: 55279 (130688f277ba) Isabelle2013 --> 2013-1: Test_Isac perfect
   223   User: Walther Neuper <neuper@ist.tugraz.at>
   224   Date: 2013-11-21 18:12:17 +0100 (2 weeks)
   225 
   226 *}
   227 
   228 subsection {* isac on Isabelle2013 *}
   229 subsubsection {* Summary of development *}
   230 text {*
   231   # Oct.13: replaced "axioms" by "axiomatization"
   232   # Oct.13: Mathias Lehnfeld started removing Unsynchornized.ref
   233   # Sep.13: integrated gcd_poly (functional, without Unsychronized.ref) into
   234     simplification of multivariate rationals (without improving the rulesets involved).
   235 *}
   236 subsubsection {* Run tests *}
   237 text {*
   238   Is standard now; this subsection will be discontinued under Isabelle2013-1
   239 *}
   240 subsubsection {* State of tests *}
   241 text {*
   242   See points in subsection "isac on Isabelle2011", "State of tests".
   243   # re-activated listC.sml
   244 *}
   245 subsubsection {* Changesets of begin and end *}
   246 text {*
   247   changeset 52174:8b055b17bd84 --- removed all code concerned with "castab = Unsynchronized.ref"
   248   User: Mathias Lehnfeld <s1210629013@students.fh-hagenberg.at>
   249   Date: Tue Nov 19 22:23:30 2013 +0000
   250   :
   251   : isac on Isablle2013 
   252   :
   253   Changeset: 52061 (4ecea2fcdc2c) --- Build_Isac.thy runs on Isabelle2013
   254   User: Walther Neuper <neuper@ist.tugraz.at>
   255   Date: 2013-07-15 08:28:50 +0200 (4 weeks)
   256 *}
   257 
   258 subsection {* isac on Isabelle2012 *}
   259 subsubsection {* Summary of development *}
   260 text {*
   261   isac on Isabelle2012 is considered just a transitional stage
   262   within the update from Isabelle2011 to Isabelle2013; thus no further development of isac;
   263   For considerations on the transition see 
   264   ~~/src/Tools/isac/Knowledge/Build_Thydata/thy, section "updating isac..".
   265 *}
   266 subsubsection {* Run tests *}
   267 text {*
   268 $ cd /usr/local/isabisac12/
   269 $ ./bin/isabelle jedit -l HOL src/Tools/isac/Build_Isac.thy
   270 $ ./bin/isabelle jedit -l Isac test/Tools/isac/Test_Isac.thy
   271 *}
   272 subsubsection {* State of tests *}
   273 text {*
   274   At least the tests from isac on Isabelle2011 run again.
   275   However, Test_Isac.thy shows erratic behaviour; no errors are obtained when scrolling 
   276   in parallel with evaluation.
   277 
   278   Counting "error in kernel" for Frontend/interface.sml (the tests considered most significant)
   279   yields 69 hits, some of which were already present before Isabelle2002-->2009-2
   280   (i.e. on the old notebook from 2002).
   281 
   282   Now many tests with (*...=== inhibit exn ...*) give a reason or at least the origin:
   283   # === inhibit exn WN1130621 Isabelle2012-->13 !thehier! === ...see Build_Thydata.thy
   284   # === inhibit exn AK110726 === ...reliable work by Alexander Kargl, most likely go back to 2002
   285   # === inhibit exn WN1130701 broken at Isabelle2002 --> 2009-2 === , most likely go back to 2002
   286   Reasons for outcommented tests are also found in Test_Isac.thy near the respective file.sml.
   287 
   288   Some tests have been re-activated (e.g. error patterns, fill patterns).
   289 *}
   290 subsubsection {* Changesets of begin and end *}
   291 text {*  
   292   Changeset: 52051 (35751d90365e) end of improving tests for isac on Isabelle2012
   293   User: Walther Neuper <neuper@ist.tugraz.at>
   294   Date: 2013-07-11 16:58:31 +0200 (4 weeks)
   295   :
   296   : isac on Isablle2012 
   297   :
   298   Changeset: 48757 (74eb3dfc33cc) updated src from Isabelle2011 to Isabelle2012
   299   User: Walther Neuper <neuper@ist.tugraz.at>
   300   Date: 2012-09-24 18:35:13 +0200 (8 months)
   301   ------------------------------------------------------------------------------
   302   Changeset: 48756 (7443906996a8) merged
   303   User: Walther Neuper <neuper@ist.tugraz.at>
   304   Date: 2012-09-24 18:15:49 +0200 (8 months)
   305 *}
   306 
   307 subsection {* isac on Isabelle2011 *}
   308 subsubsection {* Summary of development *}
   309 text {*
   310   isac's mathematics engine has been extended by two developments:
   311   (1) Isabelle's contexts were introduced by Mathias Lehnfeld
   312   (2) Z_Transform was introduced by Jan Rocnik, which revealed
   313     further errors introduced by (1).
   314   (3) "error patterns" were introduced by Gabriella Daroczy
   315   Regressions tests have been added for all of these.
   316 *}
   317 subsubsection {* Run tests *}
   318 text {*
   319   $ cd /usr/local/isabisac11/
   320   $ ./bin/isabelle jedit -l HOL src/Tools/isac/Build_Isac.thy
   321   $ ./bin/isabelle jedit -l Isac test/Tools/isac/Test_Isac.thy
   322 *}
   323 subsubsection {* State of tests *}
   324 text {* 
   325   Systematic efforts outcommented less significant tests by (*...=== inhibit exn ...*) 
   326   and sometimes give reasons for failing tests.
   327   (*...=== inhibit exn AK...*) was done by Alexander Kargl; this is reliable
   328   work, some of which couldn't be revised (and renamed) by WN and thus survived some time.
   329 
   330   The most signification tests (in particular Frontend/interface.sml) run,
   331   however, many "error in kernel" are not caught by an exception.
   332   ------------------------------------------------------------------------------
   333   After the changeset below Test_Isac worked with check_unsynchronized_ref ():
   334   ------------------------------------------------------------------------------
   335   Changeset: 42457 (ca691a84b81a) PROVISIONALLY MADE TESTS RUN with Unsynchronized.ref
   336   User: Walther Neuper <neuper@ist.tugraz.at>
   337   Date: 2012-08-06 10:38:11 +0200 (11 months)
   338 
   339 
   340   The list below records TODOs while producing an ISAC kernel for 
   341   gdaroczy and jrocnik, wich could NOT be done before all tests are RUNNING
   342   (so to be resumed with Isabelle2013-1):
   343   ############## WNxxxxxx.TODO can be found in sources ##############
   344   --------------------------------------------------------------------------------
   345   WN111013.TODO: lots of cleanup/removal in test/../Test.thy
   346   --------------------------------------------------------------------------------
   347   WN111013.TODO: remove concept around "fun init_form", lots of troubles with 
   348   this special case (see) --- why not nxt = Model_Problem here ? ---
   349   --------------------------------------------------------------------------------
   350   WN111014.TODO calculate_Poly < calculate_Rational < calculate_RootRat, see test/
   351   ... FIRST redesign 
   352   # simplify_* , *_simp_* 
   353   # norm_* 
   354   # calc_* , calculate_*  ... require iteration over all rls ...
   355   ... see --- val rls = calculate_RootRat > calculate_Rational --- CONTINUE !
   356   --------------------------------------------------------------------------------
   357   WN111014.TODO fun prep_rls | !!!use this function in ruleset' := !!!
   358   --------------------------------------------------------------------------------
   359   WN120314 changeset a393bb9f5e9f drops root equations.
   360   see test/Tools/isac/Knowledge/rootrateq.sml 
   361   --------------------------------------------------------------------------------
   362   WN120317.TODO changeset 977788dfed26 dropped rateq:
   363   # test --- repair NO asms from rls RatEq_eliminate --- shows error from 2002
   364   # test --- solve (1/x = 5, x) by me --- and --- x / (x ^ 2 - 6 * x + 9) - ...:    
   365     investigation Check_elementwise stopped due to too much effort finding out,
   366     why Check_elementwise worked in 2002 in spite of the error.
   367   --------------------------------------------------------------------------------
   368   WN120317.TODO postponed test/../ratinal,ratinal2.sml to joint work with dmeindl 
   369   --------------------------------------------------------------------------------
   370   WN120317.TODO found by test --- interSteps for Schalk 299a --- that 
   371     NO test with 'interSteps' is checked properly (with exn on changed behaviour)
   372   --------------------------------------------------------------------------------
   373   WN120317.TODO test --- Matthias Goldgruber 2003 rewrite orders --- has
   374     a newly outcommented test where rewrite_set_ make_polynomial --> NONE
   375   --------------------------------------------------------------------------------
   376   WN120320.TODO check-improve rlsthmsNOTisac:
   377   DONE make test --- old compute rlsthmsNOTisac by eq_thmI'
   378   DONE compare rlsthmsNOTisac in thms-survey-Isa02-Isa09-2.sml .. Isac.thy 
   379   FOUND 120321: Theory.axioms_of doesnt find LENGTH_CONS etc, thus are in Isab
   380   # mark twice thms (in isac + (later) in Isabelle) in Isac.thy
   381   --------------------------------------------------------------------------------
   382   WN120320.TODO rlsthmsNOTisac: replace twice thms ^
   383   --------------------------------------------------------------------------------
   384   WN120320.TODO rlsthmsNOTisac: reconsider design of sym_* thms, see test
   385   --- OLD compute rlsthmsNOTisac by eq_thmID ---: some are in isab, some in isac.
   386   --------------------------------------------------------------------------------
   387   WN120321.TODO rearrange theories:
   388     Knowledge
   389       :
   390       Atools.thy
   391       ///Descript.thy --> ProgLang
   392       Delete.thy   <--- first_Knowledge_thy (*mv to Atools.thy*)
   393     ProgLang: restructure Build_Isac.thy such that no xmlsrc.thy, ProgLang.thy
   394           Interpret.thy are generated (simplifies xml structure for theories)
   395       Script.thy
   396       Tools.thy
   397       ListC.thy    <--- first_Proglang_thy
   398   --------------------------------------------------------------------------------
   399   WN120321.TODO reanimate test/../simplify.sml: CAS-command Simplify
   400       EXN "simplify.sml: diff.behav. CAScmd: Simplify (2*a + 3*a)"
   401   broken during work on thy-hierarchy
   402   --------------------------------------------------------------------------------
   403   WN120321.TODO LAST in IsacScripts + in IsacKnowledge, see
   404   test --- the_hier (get_thes ()) (collect_thydata ())---
   405   --------------------------------------------------------------------------------
   406   WN120405a.TODO src/../pbl-met-hierarchy.sml: fun pbl2term--> Isac' instead Isac_
   407   !!add mutual crossreferences to ?fun headline??? where the same has to be done:
   408   !!OR BETTER: use only 2 functions for adding/removing "'" to/from thy!!
   409   --------------------------------------------------------------------------------
   410   WN120409.TODO replace "op mem" (2002) with member (2011) ... 
   411   ... an exercise interesting for beginners !
   412   --------------------------------------------------------------------------------
   413   WN120411 scanning html representation of newly generated knowledge:
   414   * thy:
   415   ** Theorems: only "Proof of the theorem" (correct!)
   416                and "(c) isac-team (math-autor)"
   417   ** Rulesets: only "Identifier:///"
   418                and "(c) isac-team (math-autor)"
   419   ** IsacKnowledge: link to dependency graph (which needs to be created first)
   420   ** IsacScripts --> ProgramLanguage
   421   *** Tools: Theorems: arity_type_cp, arity_type_nam, ... wegnehmen
   422   
   423   * pbl: OK !?!
   424   * met: OK !?!
   425   * exp: 
   426   ** Z-Transform is missing !!!
   427   ** type-constraints !!!
   428   --------------------------------------------------------------------------------
   429   WN120417: merging xmldata revealed:
   430   ..............NEWLY generated:........................................
   431   <THEOREMDATA>
   432     <GUH> thy_isab_Fun-thm-o_apply </GUH>
   433     <STRINGLIST>
   434       <STRING> Isabelle </STRING>
   435       <STRING> Fun </STRING>
   436       <STRING> Theorems </STRING>
   437       <STRING> o_apply </STRING>
   438     </STRINGLIST>
   439       <MATHML>
   440         <ISA> (?f o ?g) ?x = ?f (?g ?x) </ISA>
   441       </MATHML>  <PROOF>
   442       <EXTREF>
   443         <TEXT> Proof of the theorem </TEXT>
   444         <URL> http://www.ist.tugraz.at/projects/isac/www/kbase/thy/browser_info/HOL/HOL-Real/Isac/Fun.html </URL>
   445       </EXTREF>
   446     </PROOF>
   447     <EXPLANATIONS> </EXPLANATIONS>
   448     <MATHAUTHORS>
   449       <STRING> Isabelle team, TU Munich </STRING>
   450     </MATHAUTHORS>
   451     <COURSEDESIGNS>
   452     </COURSEDESIGNS>
   453   </THEOREMDATA>
   454   ..............OLD FORMAT:.............................................
   455   <THEOREMDATA>
   456     <GUH> thy_isab_Fun-thm-o_apply </GUH>
   457     <STRINGLIST>
   458       <STRING> Isabelle </STRING>
   459       <STRING> Fun </STRING>
   460       <STRING> Theorems </STRING>
   461       <STRING> o_apply </STRING>
   462     </STRINGLIST>
   463     <THEOREM>
   464       <ID> o_apply </ID>
   465       <MATHML>
   466         <ISA> (?f o ?g) ?x = ?f (?g ?x) </ISA>
   467       </MATHML>
   468     </THEOREM>
   469     <PROOF>
   470       <EXTREF>
   471         <TEXT> Proof of the theorem </TEXT>
   472         <URL> http://www.ist.tugraz.at/projects/isac/www/kbase/thy/browser_info/HOL/HOL-Real/Isac/Fun.html </URL>
   473       </EXTREF>
   474     </PROOF>
   475     <EXPLANATIONS> </EXPLANATIONS>
   476     <MATHAUTHORS>
   477       <STRING> Isabelle team, TU Munich </STRING>
   478     </MATHAUTHORS>
   479     <COURSEDESIGNS>
   480     </COURSEDESIGNS>
   481   </THEOREMDATA>
   482   --------------------------------------------------------------------------------
   483 *}
   484 subsubsection {* Changesets of begin and end *}
   485 text {*
   486   isac development was done between these changesets:
   487   ------------------------------------------------------------------------------
   488   Changeset: 42519 (1f3b4270363e) meeting dmeindl: added missing files
   489   User: Walther Neuper <neuper@ist.tugraz.at>
   490   Date: 2012-09-24 16:39:30 +0200 (8 months)
   491   :
   492   : isac on Isablle2011
   493   :
   494   Changeset:41897 (355be7f60389) merged isabisac with Isabelle2011
   495   Branch: decompose-isar 
   496   User: Walther Neuper <neuper@ist.tugraz.at>
   497   Date: 2011-02-25 13:04:56 +0100 (2011-02-25)
   498   ------------------------------------------------------------------------------
   499 *}
   500 
   501 subsection {* isac on Isabelle2009-2 *}
   502 subsubsection {* Summary of development *}
   503 text {*
   504   In 2009 the update of isac from Isabelle2002 started with switching from CVS to hg.
   505   The update was painful (bridging 7 years of Isabelle development) and cut short 
   506   due to the start of introducing Isabelle's contexts (Mathias Lehnfeld) and
   507   going on to Isabelle2011 although most of the tests did not run.
   508 *}
   509 subsubsection {* Run tests *}
   510 text {*
   511   WN131021 this is broken by installation of Isabelle2011/12/13,
   512   because all these write their binaries to ~/.isabelle/heaps/..
   513 
   514   $ cd /usr/local/isabisac09-2/
   515   $ ./bin/isabelle emacs -l HOL src/Tools/isac/Build_Isac.thy
   516   $ ./bin/isabelle emacs -l Isac src/Tools/isac/Test_Isac.thy
   517   NOT THE RIGHT VERSION.....    test/Tools/isac/Test_Isac.thy !!!
   518 *}
   519 subsubsection {* State of tests *}
   520 text {* 
   521   Most tests are broken by the update from Isabelle2002 to Isabelle2009-2.
   522 *}
   523 subsubsection {* Changesets of begin and end *}
   524 text {*
   525   isac development was done between these changesets:
   526   ------------------------------------------------------------------------------
   527   Changeset: 38115 (940a5feea094) Little improvements of isac-plugin
   528   Branch: decompose-isar 
   529   User: Marco Steger <m.steger@student.tugraz.at>
   530   Date: 2011-02-06 18:30:28 +0100 (2011-02-06)
   531   :
   532   : isac on Isablle2009-2
   533   :
   534   Changeset: 37870 (5100a9c3abf8) created branch isac-from-Isabelle2009-2
   535   Branch: isac-from-Isabelle2009-2 
   536   User: Walther Neuper <neuper@ist.tugraz.at>
   537   Date: 2010-07-21 09:59:35 +0200 (2010-07-21)
   538   ------------------------------------------------------------------------------
   539 *}
   540 
   541 subsection {* isac on Isabelle2002 *}
   542 subsubsection {* Summary of development *}
   543 text {*
   544   From 1999 to 2010 all the basic functionality (except "ctxt" and "error pattern")
   545   of isac's mathematics engine has been implemented.
   546 *}
   547 subsubsection {* Run tests *}
   548 subsubsection {* State of tests *}
   549 text {* 
   550   All tests work on an old notebook (the right PolyML coudn't be upgraded to more
   551   recent Linux versions)
   552 *}
   553 subsubsection {* Changesets of begin and end *}
   554 text {*
   555   Up to the above Mercurial changeset 5100a9c3abf8 isac used CVS;
   556   see https://intra.ist.tugraz.at/hg/isac containing a conversion to Mercurial.
   557 *}
   558 
   559 end
   560 (*========== inhibit exn 130719 Isabelle2013 ===================================
   561 ============ inhibit exn 130719 Isabelle2013 =================================*)
   562 
   563 (*-.-.-.-.-.-isolate response.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.
   564   -.-.-.-.-.-.-isolate response.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.-.*)
   565