src/Tools/isac/TODO.thy
author Walther Neuper <walther.neuper@jku.at>
Thu, 14 May 2020 13:33:47 +0200
changeset 59977 e635534c5f63
parent 59976 950922a768ca
child 59981 dc34eff67648
permissions -rw-r--r--
rename Specification -> References, contiued
     1 (* Title:  todo's for isac core
     2    Author: Walther Neuper 111013
     3    (c) copyright due to lincense terms.
     4 *)
     5 theory TODO
     6 imports "~~/src/Doc/Prog_Prove/LaTeXsugar"
     7 begin
     8 
     9 text \<open>
    10 Legend: here + code since 200120
    11   \<open>(*DOC\<close> identifies stuff dedicated for isac/Doc/Lucas_Interpreter
    12   \<open>(*?!?\<close> identifies a design question
    13 
    14 Distillation of "rounds of reforms":
    15   \<open>TODOs from current changesets\<close> \<rightarrow> \<open>Postponed from current changeset\<close> \<rightarrow>
    16   (*to be removed..*) \<rightarrow> \<open>Separated tasks\<close> (*.. to be removed*)
    17   \<rightarrow> subsection of \<open>Major reforms\<close>
    18 \<close>
    19 
    20 section \<open>TODOs from current changesets\<close>
    21 text\<open>
    22   Shift more complicated issues from  \<open>Current changeset\<close> to \<open>Postponed from current changeset\<close>
    23 \<close>
    24 subsection \<open>Current changeset\<close>
    25 text \<open>
    26 (*/------- to  from -------\*)
    27 (*\------- to  from -------/*)
    28   \begin{itemize}
    29   \item xxx
    30   \item rename Tactic.Calculate -> Tactic.Evaluate
    31   \item xxx
    32   \item replace src/ Erls by Rule_Set.Empty
    33   \item xxx
    34   \item rename ptyps.sml -> specify-etc.sml
    35         rename Specify -> Specify_Etc
    36         rename Specify -> Specify
    37   \item xxx
    38   \item specific_from_prog in..end: replace o by |> (Test_Isac or Test_Some!)
    39   \item xxx
    40   \item cleanup ThmC in MathEngBasic
    41   \item xxx
    42   \item xxx
    43   \item xxx
    44   \item xxx
    45   \item xxx
    46   \item xxx
    47   \item xxx
    48   \item xxx
    49   \end{itemize}
    50 \<close>
    51 subsection \<open>Postponed from current changeset\<close>
    52 text \<open>
    53   \begin{itemize}
    54   \item xxx
    55   \item revise O_Model and I_Model with an example with more than 1 variant.
    56   \item xxx
    57   \item ctxt is superfluous in specify-phase due to type constraints from Descript.thy
    58   \item xxx
    59   \item Derive.do_one: TODO find code in common with complete_solve
    60         Derive.embed: TODO rewrite according to CAO (+ no intermediate access to Ctree)
    61   \item xxx
    62   \item Solve_Check: postponed parsing input to _ option
    63   \item xxx
    64   \item ? "fetch-tactics.sml" from Mathengine -> BridgeLibisabelle ?
    65   \item xxx
    66   \item ? unify struct.Step and struct.Solve in MathEngine ?
    67   \item xxx
    68   \item use "Eval_Def" for renaming identifiers
    69   \item xxx
    70   \item why does Test_Build_Thydata.thy depend on ProgLang and not on CalcElements ?
    71   \item xxx
    72   \item xxx
    73   \begin{itemize}
    74   \item LI.do_next (*TODO RM..*) ???
    75   \item generate.sml: RM datatype edit TOGETHER WITH datatype inout
    76   \item TermC.list2isalist: typ -> term list -> term  .. should NOT requires typ
    77   \item get_ctxt_LI should replace get_ctxt
    78   \item ALL CODE: rename spec(ification) --> know(ledge), in Specification: Relation -> Knowledge
    79   \item rename   Base_Tool.thy  <---         Base_Tools
    80   \item adopt naming conventions in Knowledge: EqSystem --> Equation_System, etc
    81   \item rename field scr in meth
    82   \item DEL double code: nxt_specify_init_calc IN specify.sml + step-specify.sml
    83   \item xxx
    84   \item xxx
    85   \item xxx
    86   \item xxx
    87   \item xxx
    88   \item xxx
    89   \item xxx
    90   \item clarify Tactic.Subproblem (domID, pblID) as term in Pstate {act_arg, ...}
    91      there it is Free ("Subproblem", "char list \<times> ..
    92      instead of  Const (|???.Subproblem", "char list \<times> ..
    93      AND THE STRING REPRESENTATION IS STRANGE: 
    94        Subproblem (''Test'',
    95         ??.\ <^const> String.char.Char ''LINEAR'' ''univariate'' ''equation''
    96          ''test'')
    97      ML \<open>
    98      \<close> ML \<open>
    99      term2str; (*defined by..*)
   100      fun term_to_string''' thy t =
   101        let
   102          val ctxt' = Config.put show_markup false (Proof_Context.init_global thy)
   103        in Print_Mode.setmp [] (Syntax.string_of_term ctxt') t end;
   104      \<close> ML \<open>
   105      val t = @{term "aaa + bbb"}
   106      val t = @{term "Subproblem (''Test'', [''LINEAR'', ''univariate'', ''equation''])"};
   107      \<close> ML \<open>
   108      val sss = Print_Mode.setmp [] (Syntax.string_of_term @{context}) t
   109      \<close> ML \<open>
   110      writeln sss (*.. here perfect: Subproblem (''Test'', [''LINEAR'', ''univariate'', ''equation'']) *)
   111      \<close> ML \<open>
   112      \<close>
   113   \item xxx
   114   \item xxx
   115   \item cleanup fun me:
   116     fun me (*(_, Empty_Tac) p _ _ = raise ERROR ("me:  Empty_Tac at " ^ pos'2str p)
   117       | me*) (_, tac) p _(*NEW remove*) pt =
   118     + -------------------------^^^^^^
   119     # see test-code.sml fun me_trace
   120       use this also for me, not only for me_ist_ctxt; del. me
   121       this requires much updating in all test/*
   122   \item xxx
   123   \item shift tests into NEW model.sml (upd, upds_envv, ..)
   124   \item xxx
   125   \item clarify handling of contexts ctxt ContextC
   126     \begin{itemize}
   127     \item xxx
   128     \item Specify/ works with thy | Interpret/ works with ctxt | MathEngine.step works with ?!?ctxt
   129     \item xxx
   130     \item Check_Elementwise "Assumptions": prerequisite for ^^goal
   131       rm tactic Check_elementwise "Assumptions" in a way, which keeps it for Minisubpbl
   132       rm Assumptions  :: bool  (* TODO: remove with making ^^^ idle *)
   133     \item xxx
   134     \item xxx
   135     \end{itemize}
   136   \item xxx
   137   \item xxx
   138   \item complete mstools.sml (* survey on handling contexts:
   139   \item xxx
   140   \item xxx
   141   \item xxx
   142   \item xxx
   143   \item librarys.ml --> libraryC.sml + text from termC.sml
   144   \item xxx
   145   \item xxx
   146   \item xxx
   147   \item xxx
   148   \item xxx
   149   \item xxx
   150   \item concentrate "insert_assumptions" for locate_input_tactic in "associate", ?OR? Tactic.insert_assumptions
   151                                    DONE  for find_next_step by Tactic.insert_assumptions m' ctxt 
   152     \begin{itemize}
   153     \item rm from "generate1" ("Detail_Set_Inst'", Tactic.Detail_Set' ?)
   154     \item ?"insert_assumptions" necessary in "init_pstate" ?+++? in "applicable_in" ?+++? "associate"
   155     \item xxx
   156     \item xxx
   157     \item DO DELETIONS AFTER analogous concentrations in find_next_step
   158     \end{itemize}
   159   \item xxx
   160   \item ? what is the difference headline <--> cascmd in
   161     Subproblem' (spec, oris, headline, fmz_, context, cascmd)
   162   \item Substitute' what does "re-calculation mean in the datatype comment?
   163   \item xxx
   164   \item inform: TermC.parse (ThyC.get_theory "Isac_Knowledge") istr --> parseNEW context istr
   165   \item xxx
   166   \item unify/clarify stac2tac_ | 
   167     \begin{itemize}
   168     \item xxx
   169     \item extract common code from associate.. stac2tac_xxx
   170     \item rename LItool.tac_from_prog -> Tactic.from_prog_tac ? Solve_Tac.from_prog_tac,
   171     \item xxx
   172     \item xxx
   173     \end{itemize}
   174   \item xxx
   175   \item unify in signature LANGUAGE_TOOLS =\\
   176     val pblterm: ThyC.id -> Problem.id -> term     vvv        vvv\\
   177     val subpbl: string -> string list -> term          unify with ^^^
   178   \item xxx
   179   \item Telem.safe is questionable: has it been replaced by Safe_Step, Not_Derivable, Helpless, etc?
   180     Note: replacement of Istate.safe by Istate.appy_ didn't care much about Telem.safe.
   181     If Telem.safe is kept, consider merge with CTbasic.ostate
   182   \item xxx
   183   \item remove find_next_step from solve Apply_Method';
   184     this enforces Pos.at_first_tactic, which should be dropped, too.
   185   \item xxx
   186   \item xxx
   187   \item xxx
   188   \end{itemize}
   189 \<close>
   190 subsection \<open>Postponed --> Major reforms\<close>
   191 text \<open>
   192   \begin{itemize}
   193   \item xxx          
   194   \item revisit bootstrap Calcelements. rule->calcelems->termC
   195     would be nice, but is hard: UnparseC.terms -> TermC.s_to_string
   196   \item xxx
   197   \item replace all Ctree.update_* with Ctree.cupdate_problem
   198   \item xxx
   199   \item rename (ist as {eval, ...}) -> (ist as {eval_rls, ...})
   200   \item xxx
   201   \item exception PROG analogous to TERM
   202   \item xxx
   203   \item sig/struc ..elems --> ..elem
   204   \item xxx          
   205   \item distille CAS-command, CAScmd, etc into a struct
   206   \item xxx          
   207   \item check location of files:
   208         test/Tools/isac/Interpret/ptyps.thy
   209         test/Tools/isac/Specify.ptyps.sml
   210   \item xxx
   211   \item check occurences of Atools in src/ test/
   212   \item Const ("Atools.pow", _) ---> Const ("Base_Tool.pow", _) 
   213   \item xxx          
   214   \item rm Float
   215   \item xxx
   216   \item Diff.thy: differentiateX --> differentiate after removal of script-constant
   217   \item Test.thy: met_test_sqrt2: deleted?!
   218   \item xxx          
   219   \item Rewrite_Ord.rew_ord' := overwritel (! Rewrite_Ord.rew_ord', (*<<<---- use Know_Store.xxx, too*)
   220   \item xxx
   221     \item automatically extrac rls from program-code 
   222       ? take ["SignalProcessing", "Z_Transform", "Inverse_sub"] as an example ?
   223   \item xxx          
   224   \item finish output of LItool.trace with Check_Postcond (useful for SubProblem)
   225   \item xxx
   226   \item replace Rule_Set.empty by Rule_Set.Empty
   227     latter is more clear, but replacing ***breaks rewriting over all examples***,
   228     e.g. see ERROR: rewrite__set_ called with 'Erls' for 'precond_rootpbl x'
   229     in Minisubplb/200-start-method-NEXT_STEP.sml:
   230      (*+* )------- in f3cac3053e7b (Rule_Set.empty just renamed, NOT deleted) we had
   231      (*+*)  prls =
   232      (*+*)    Rls {calc = [], erls = Erls, errpatts = [], id = "empty", preconds = [], rew_ord =
   233      (*+*)      ("dummy_ord", fn), rules = [], scr = Empty_Prog, srls = Erls}:
   234      (*+*).. THIS IS Rule_Set.empty, BUT IT DID not CAUSE ANY ERROR !
   235      (*+*)------- WITH Rule_Set.empty REMOVED (based on f3cac3053e7b) we had
   236      (*+*)val Empty = prls (* <---ERROR: rewrite__set_ called with 'Erls' for 'precond_rootpbl x' *)
   237      ( *+*)val ["sqroot-test", "univariate", "equation", "test"] = cpI
   238     THAT INDICATES, that much rewriting/evaluating JUST WORKED BY CHANCE?!?
   239   \item xxx          
   240   \item xxx
   241   \item xxx          
   242   \item xxx
   243   \item xxx          
   244   \item xxx
   245   \item xxx          
   246   \end{itemize}
   247 \<close>
   248 
   249 section \<open>Major reforms\<close>
   250 text \<open>
   251 \<close>
   252 subsection \<open>Exception Size raised\<close>
   253 text \<open>
   254   During update Isabelle2018 --> Isabelle2019 we noticed, that
   255   "isabelle build" uses resources more efficiently than "isabelle jedit".
   256   The former works, but the latter causes 
   257   \begin{itemize}
   258   \item "Exception- Size raised"
   259     in Build_Thydata.thy
   260   \item "exception Size raised (line 169 of "./basis/LibrarySupport.sml")"
   261     in test/../biegelinie-*.xml.
   262   \end{itemize}
   263   This has been detected after changeset (30cd47104ad7) "lucin: reorganise theories in ProgLang".
   264 
   265   Find tools to investigate the Exception, and find ways around it eventually.
   266 \<close>
   267 subsection \<open>Cleanup & review signatures wrt. implementation.pdf canonical argument order\<close>
   268 text \<open>
   269   \begin{itemize}
   270   \item there are comments in several signatures
   271   \item ML_file "~~/src/Tools/isac/Interpret/specification-elems.sml" can be (almost) deleted
   272   \item src/../Frontend/: signatures missing
   273   \item xxx
   274   \end{itemize}
   275 \<close>
   276 subsection \<open>overall structure of code\<close>
   277 text \<open>
   278   \begin{itemize}
   279   \item try to separate Isac_Knowledge from MathEngine
   280     common base: Knowledge_Author / ..??
   281   \item xxx
   282   \item  ML_file "~~/src/Tools/isac/Interpret/ctree.sml" (*shift to base in common with Interpret*)
   283   \item xxx
   284   \item xxx
   285   \item xxx
   286   \end{itemize}
   287 \<close>
   288 subsection \<open>Separate MathEngineBasic/ Specify/ Interpret/ MathEngine/\<close>
   289 text \<open>
   290   \begin{itemize}
   291   \item xxx
   292   \item xxx
   293   \item re-organise code for Interpret
   294     \begin{itemize}
   295     \item Step*: Step_Specify | Step_Solve | Step DONE
   296     \item xxx
   297     \item Prog_Tac: fun get_first_argument takes both Prog_Tac + Program --- wait for separate Tactical
   298       then shift into common descendant
   299     \item xxx
   300     \end{itemize}
   301   \item xxx
   302   \item xxx
   303   \item ??????????? WHY CAN LI.by_tactic NOT BE REPLACED BY Step_Solve.by_tactic ???????????
   304   \item xxx
   305   \item xxx
   306   \end{itemize}
   307 \<close>
   308 subsection \<open>Review modelling- + specification-phase\<close>
   309 text \<open>
   310   \begin{itemize}
   311   \item xxx
   312   \item xxx
   313   \item xxx
   314   \item check match between args of partial_function and model-pattern of meth;
   315      provide error message.
   316   \item xxx
   317   \item "--- hack for funpack: generalise handling of meths which extend problem items ---"
   318     \begin{itemize}
   319     \item see several locations of hack in code
   320     \item these locations are NOT sufficient, see
   321       test/../biegelinie-3.sml --- IntegrierenUndKonstanteBestimmen2: Bsp.7.70. auto ---
   322     \item "fun associate" "match_ags ..dI" instead "..pI" breaks many tests, however,
   323       this would be according to survey Notes (3) in src/../calchead.sml.
   324     \end{itemize}
   325   \item see "failed trial to generalise handling of meths"98298342fb6d
   326   \item abstract specify + nxt_specif to common aux-funs;
   327     see e.g. "--- hack for funpack: generalise handling of meths which extend problem items ---"
   328   \item xxx
   329   \item type model = itm list ?
   330   \item review survey Notes in src/../calchead.sml: they are questionable
   331   \item review copy-named, probably two issues commingled 
   332     \begin{itemize}
   333     \item special handling of "#Find#, because it is not a formal argument of partial_function
   334     \item special naming for solutions of equation solving: x_1, x_2, ...
   335     \end{itemize}
   336   \item xxx
   337   \item xxx
   338   \item this has been written in one go:
   339     \begin{itemize}
   340     \item reconsidering I_Model.max_vt, use problem with meth ["DiffApp","max_by_calculus"]
   341     \item reconsider add_field': where is it used for what? Shift into mk_oris
   342     \item reconsider match_itms_oris: where is it used for what? max_vt ONLY???
   343     \item in Specify_Method search root-oris for items (e.g. "errorBound"), #1# in survey
   344     \item Specify_Problem, Specify_Method: check respective identifiers after re-Specify_
   345       (relevant for pre-condition)
   346     \item unify match_ags to mk_oris1..N with different args (fmz | pat list, pbl | meth
   347     \item 
   348     \end{itemize}
   349   \end{itemize}
   350 \<close>
   351 subsection \<open>taci list, type step\<close>
   352 text \<open>
   353 taci was, most likely, invented to make "fun me" more efficient by avoiding duplicate rewrite,
   354 and probably, because the Kernel interface separated setNextTactic and applyTactic.
   355 Both are no good reasons to make code more complicated.
   356 
   357 !!! taci list is used in do_next !!!
   358 
   359   \begin{itemize}
   360   \item xxx
   361   \item can lev_on_total replace lev_on ? ..Test_Isac_Short + rename lev_on_total -> lev_on
   362   \item xxx
   363   \item Step* functions should return Calc.T instead of Specification.calcstate'
   364   \item xxx
   365   \item states.sml: check, when "length tacis > 1"
   366   \item in Test_Isac.thy there is only 1 error in Interpret/inform.sml
   367   \item (*WN190713 REMOVE: "creating a new node" was never implemented for more than one node?!?
   368   \item xxx
   369   \item brute force setting all empty ([], [], ptp) works!?! but ptp causes errors -- investigate!
   370   \item xxx
   371   \end{itemize}
   372 \<close>
   373 subsection \<open>Ctree\<close>
   374 text \<open>
   375 analysis
   376 # mixture pos' .. pos in cappend_*, append_* is confusing
   377 # existpt p pt andalso Tactic.is_empty DIFFERENT IN append_*, cappend_* is confusing
   378   "exception PTREE "get_obj: pos =" ^^^^^: ^^^^ due to cut !!!
   379   NOTE: exn IN if..andalso.. IS NOT!!! DETECTED, THIS                       is confusing
   380   see test/../--- Minisubpbl/800-append-on-Frm.sml ---
   381 # ?!? "cut branches below cannot be decided here" in append_atomic
   382 # sign. of functions too different ?!?canonical arg.order ?!?
   383   \begin{itemize}
   384   \item xxx
   385   \item remove update_branch, update_*? -- make branch, etc args of append_*
   386   \item xxx
   387   \item close sig Ctree, contains cappend_* ?only? --- ?make parallel to ?Pide_Store?
   388   \item xxx
   389   \item unify args to Ctree.state (pt, p)
   390   \item  fun update_env       .. repl_env                   \<rightarrow>updatempty
   391   \item xxx
   392   \item xxx
   393   \item xxx
   394   \item xxx
   395   \end{itemize}
   396 \<close>
   397 subsection \<open>replace theory/thy by context/ctxt\<close>
   398 text \<open>
   399   \begin{itemize}
   400   \item xxx
   401   \item Specify/ works with thy | Interpret/ works with ctxt | MathEngine.step works with ?!?ctxt
   402     special case: Tactic.Refine_Problem
   403   \item xxx
   404   \item theory can be retrieved from ctxt by Proof_Context.theory_of
   405   \item xxx
   406   \item cleaup the many conversions string -- theory
   407   \item make dest_spec --> (theory, pblID, metID) ?+ common_subthy ?
   408   \item 1. Rewrite.eval_true_, then
   409     LItool.check_leaf, Rewrite.eval_prog_expr, Step.add, LItool.tac_from_prog.
   410   \item fun associate
   411     let val thy = ThyC.get_theory "Isac_Knowledge";(*TODO*)
   412   \item xxx
   413   \item xxx
   414   \item in locate_input_tactic .. ?scan_dn1?; Program.is_eval_expr   .use  Term.exists_Const
   415   \item push srls into pstate
   416   \item lucas-intrpreter.locate_input_tactic: scan_to_tactic1 srls tac cstate (progr, Rule_Set.Empty)
   417                                                                                       ^^^^^^^^^^^^^^^
   418   \item xxx
   419   \end{itemize}
   420 \<close>
   421 subsection \<open>Rfuns, Begin_/End_Detail', Rrls, Istate\<close>
   422 text \<open>
   423 remove refactor Rfuns, Rule.Prog, Rule.Empty_Prog, RrlsState: this is a concept never brought to work.
   424   Clarify relation to reverse rewriting!
   425   \begin{itemize}
   426   \item separate mut.recursion program with rule and rls by deleting fild scr in rls
   427     (possible since CS 43160c1e775a
   428   ` "replace Prog. in prep_rls by Auto_Prog.gen, which generates Prog. on the fly" )
   429   \item xxx
   430   \item probably only "normal_form" seems to be needed
   431   \item deleted Rfuns in NEW "locate_input_tactic": no active test for "locate_rule"
   432     but that seems desirable
   433   \item ?how is the relation to reverse-rewriting ???
   434   \item "Rfuns" markers in test/../rational
   435   \item xxx
   436   \item datatype istate (Istate.T): remove RrlsState, pstate: use Rrls only for creating results beyond
   437     rewriting and/or respective intermediate steps (e.g. cancellation of fractions).
   438     Thus we get a 1-step-action which does NOT require a state beyond istate/pstate.
   439     Thus we drastically reduce complexity, also get rid of "fun from_pblobj_or_detail_calc" , etc.
   440   \item debug ^^^ related: is there an occurence of Steps with more than 1 element?
   441   \item xxx
   442   \item and do_next (* WN1907: ?only for Begin_/End_Detail' DEL!!!*)
   443   \item xxx
   444   \item shouldn't go Rfuns from Rewrite --> Rewrite_Set; they behave similar to "fun interSteps" ?
   445   \item xxx
   446   \item ?finally Prog could go from Calcelems to ProgLang?
   447   \end{itemize}
   448 \<close>
   449 subsection \<open>Inverse_Z_Transform.thy\<close>
   450 text \<open>
   451   \begin{itemize}
   452   \item\label{new-var-rhs} rule1..6, ruleZY introduce new variables on the rhs of the rewrite-rule.
   453   ? replace rewriting with substitution ?!?
   454   The problem is related to the decision of typing for "d_d" and making bound variables free (while
   455   shifting specific handling in equation solving etc. to the meta-logic). 
   456   \item Find "stepResponse (x[n::real]::bool)" is superfluous, because immediately used by
   457     rewrite-rules; see \ref{new-var-rhs}.
   458   \item Reconsider whole problem:
   459     input only the polynomial as argument of partial_function, in ([1], Frm) compile lhs "X z" ?
   460   \end{itemize}
   461 \<close>
   462 subsection \<open>Adopt Isabelle's numerals for Isac\<close>
   463 text \<open>
   464   \begin{itemize}
   465   \item replace numerals of type "real" by "nat" in some specific functions from ListC.thy
   466     and have both representations in parallel for "nat".
   467   \item xxx
   468   \item xxx
   469   \end{itemize}
   470 \<close>
   471 subsection \<open>Redesign equation solver\<close>
   472 text \<open>
   473   Existing solver is structured along the WRONG assumption,
   474   that Poly.thy must be the LAST thy among all thys involved -- while the opposite is the case.
   475 
   476   Preliminary solution: all inappropriately located thms are collected in Base_Tools.thy
   477 \<close>
   478 subsection \<open>Finetunig required for xmldata in kbase\<close>
   479 text \<open>
   480   See xmldata https://intra.ist.tugraz.at/hg/isac/rev/5b222a649390
   481   and in kbase html-representation generated from these xmldata.
   482   Notes in ~~/xmldata/TODO.txt.
   483 \<close>
   484 
   485 section \<open>Hints for further development\<close>
   486 text \<open>
   487 \<close>
   488 subsection \<open>Coding standards & some explanations for math-authors\<close>
   489 text \<open>copy from doc/math-eng.tex WN.28.3.03
   490 WN071228 extended\<close>
   491 
   492 subsubsection \<open>Identifiers\<close>
   493 text \<open>Naming is particularily crucial, because Isabelles name space is global, and isac does
   494   not yet use the novel locale features introduces by Isar. For instance, {\tt probe} sounds
   495   reasonable as (1) a description in the model of a problem-pattern, (2) as an element of the
   496   problem hierarchies key, (3) as a socalled CAS-command, (4) as the name of a related script etc.
   497   However, all the cases (1)..(4) require different typing for one and the same
   498   identifier {\tt probe} which is impossible, and actually leads to strange errors
   499   (for instance (1) is used as string, except in a script addressing a Subproblem).
   500 
   501   These are the preliminary rules for naming identifiers>
   502   \begin{description}
   503   \item [elements of a key] into the hierarchy of problems or methods must not contain
   504     capital letters and may contain underscrores, e.g. {\tt probe, for_polynomials}.
   505   \item [descriptions in problem-patterns] must contain at least 1 capital letter and
   506     must not contain underscores, e.g. {\tt Probe, forPolynomials}.
   507   \item [CAS-commands] follow the same rules as descriptions in problem-patterns above, thus
   508     beware of conflicts~!
   509   \item [script identifiers] always end with {\tt Program}, e.g. {\tt ProbeScript}.
   510   \item [???] ???
   511   \item [???] ???
   512   \end{description}
   513 %WN071228 extended\<close>
   514 
   515 subsubsection \<open>Rule sets\<close>
   516 text \<open>The actual version of the coding standards for rulesets is in {\tt /IsacKnowledge/Atools.ML
   517   where it can be viewed using the knowledge browsers.
   518 
   519   There are rulesets visible to the student, and there are rulesets visible (in general) only for
   520   math authors. There are also rulesets which {\em must} exist for {\em each} theory;
   521   these contain the identifier of the respective theory (including all capital letters) 
   522   as indicated by {\it Thy} below.
   523 
   524   \begin{description}
   525   \item [norm\_{\it Thy}] exists for each theory, and {\em efficiently} calculates a
   526     normalform for all terms which can be expressed by the definitions of the respective theory
   527     (and the respective parents).
   528   \item [simplify\_{\it Thy}] exists for each theory, and calculates a normalform for all terms
   529     which can be expressed by the definitions of the respective theory (and the respective parents)
   530     such, that the rewrites can be presented to the student.
   531   \item [calculate\_{\it Thy}] exists for each theory, and evaluates terms with 
   532     numerical constants only (i.e. all terms which can be expressed by the definitions of
   533     the respective theory and the respective parent theories). In particular, this ruleset includes
   534     evaluating in/equalities with numerical constants only.
   535     WN.3.7.03: may be dropped due to more generality: numericals and non-numericals
   536     are logically equivalent, where the latter often add to the assumptions
   537     (e.g. in Check_elementwise).
   538   \end{description}
   539 
   540   The above rulesets are all visible to the user, and also may be input; 
   541   thus they must be contained in {\tt Theory_Data} (KEStore_Elems.add_rlss,
   542   KEStore_Elems.get_rlss). All these rulesets must undergo a preparation
   543   using the function {\tt prep_rls'}, which generates a script for stepwise rewriting etc.
   544   The following rulesets are used for internal purposes and usually invisible to the (naive) user:
   545 
   546   \begin{description}
   547   \item [*\_erls] TODO
   548   \item [*\_prls] 
   549   \item [*\_srls] 
   550   \end{description}
   551 {\tt Rule_Set.append_rules, Rule_Set.merge, remove_rls} TODO
   552 \<close>
   553 
   554 subsection \<open>get proof-state\<close>
   555 text \<open>
   556   Re: [isabelle] Programatically get "this"
   557   ----------------------------------------------------
   558   So here is my (Makarius') version of your initial example, following these principles:
   559   begin{verbatim}
   560     notepad
   561     begin
   562       assume a: A
   563       ML_val \<open>
   564         val ctxt = @{context};
   565     
   566         val this_name =
   567           Name_Space.full_name (Proof_Context.naming_of ctxt) (Binding.name Auto_Bind.thisN);
   568         val this = #thms (the (Proof_Context.lookup_fact ctxt this_name));
   569       \<close>
   570     end
   571   end{verbatim}
   572 \<close>
   573 subsection \<open>write Specification to jEdit\<close>
   574 text \<open>
   575   Re: [isabelle] Printing terms with type annotations
   576   ----------------------------------------------------
   577   On 06/02/2019 17:52, Moa Johansson wrote:
   578   >
   579   > I’m writing some code that should create a snippet of Isar script.
   580   
   581   This is how Sledgehammer approximates this:
   582   
   583   http://isabelle.in.tum.de/repos/isabelle/file/Isabelle2018/src/HOL/Tools/Sledgehammer/sledgehammer_isar_proof.ML#l299
   584   
   585   (The module name already shows that the proper terminology is "Isar
   586   proof" (or "Isar proof text").  Proof scripts are a thing from the past,
   587   before Isar. You can emulate old-style proof scripts via a sequence of
   588   'apply' commands, but this is improper Isar.)
   589   
   590   Note that there is no standard function in Isabelle/Pure, because the
   591   problem to print just the right amount of type information is very
   592   complex and not fully solved. One day, after 1 or 2 rounds of
   593   refinements over the above approach, it might become generally available.
   594 \<close>
   595 subsection \<open>follow Isabelle conventions (*Does not yet work in Isabelle2018\<close>
   596 text \<open>
   597   isabelle update -u path_cartouches
   598   isabelle update -u inner_syntax_cartouches
   599 \<close>
   600 section \<open>Questions to Isabelle experts\<close>
   601 text \<open>
   602 \begin{itemize}
   603 \item ad ERROR Undefined fact "all_left"        in Test_Isac: error-pattern.sml
   604                Undefined fact: "xfoldr_Nil"                   inssort.sml
   605     (* probably two different reasons:
   606     src/../LinEq.thy
   607       (*WN0509 compare PolyEq.all_left "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)"*)
   608       all_left:          "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)" and
   609     
   610     src/../PolyEq.thy
   611       (*WN0509 compare LinEq.all_left "[|Not(b=!=0)|] ==> (a=b) = (a+(-1)*b=0)"*)
   612       all_left:              "[|Not(b=!=0)|] ==> (a = b) = (a - b = 0)" and
   613     
   614     test/../partial_fractions.sml
   615     (*[7], Met*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Apply_Method ["PolyEq", "normalise_poly"])*)
   616     (*[7, 1], Frm*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt = Rewrite ("all_left", "\<not> ?b =!= 0 \<Longrightarrow> (?a = ?b) = (?a - ?b = 0)"))*)
   617     
   618     test/../mathengine-stateless.sml
   619     (*if ThmC.string_of_thm @ {thm rnorm_equation_add} =  "\<not> ?b =!= 0 \<Longrightarrow> (?a = ?b) = (?a + - 1 * ?b = 0)"
   620     then () else error "string_of_thm changed";*)
   621     
   622     (*---------------------------------------------------------------------------------------------*)
   623     src/../LinEq.thy
   624     primrec xfoldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a xlist \<Rightarrow> 'b \<Rightarrow> 'b" where
   625     xfoldr_Nil:  "xfoldr f {|| ||} = id" |
   626     xfoldr_Cons: "xfoldr f (x @# xs) = f x \<circ> xfoldr f xs"
   627     
   628     src/../InsSort.thy
   629         srls = Rule_Set.empty, calc = [], rules = [
   630           Rule.Thm ("xfoldr_Nil",(*num_str*) @{thm xfoldr_Nil} (* foldr ?f [] = id *)),
   631     *)
   632 \item xxx
   633 \item xxx
   634 \item ?OK Test_Isac_Short with
   635     LI.by_tactic tac (get_istate_LI pt p, get_ctxt_LI pt p) ptp
   636   instead
   637     LI.by_tactic tac (Istate.empty, ContextC.empty) ptp
   638   in step-solve ?
   639 \item xxx
   640 \item test from last CS with outcommented re-def of code -> 
   641   -> \<open>further tests additional to src/.. files\<close>
   642       ADDTESTS/redefined-code.sml
   643 \item xxx
   644 \item efb749b79361 Test_Some_Short.thy has 2 errors, which disappear in thy ?!?:
   645   ML_file "Interpret/error-pattern.sml" Undefined fact: "all_left"
   646   ML_file "Knowledge/inssort.sml" Undefined fact: "xfoldr_Nil"
   647 \item xxx
   648 \item what is the actual replacement of "hg log --follow" ?
   649 \item xxx
   650 \item how HANDLE these exceptions, e.g.:
   651     Syntax.read_term ctxt "Randbedingungen y 0 = (0::real), y L = 0, M_b 0 = 0, M_b L = 0]"
   652   GIVES
   653     "Inner syntax error
   654      Failed to parse term"
   655 \item xxx
   656 \item how cope with "exception Size raised (line 171 of "./basis/LibrarySupport.sml")"
   657   e.g. in test/Interpret/lucas-interpreter.sml
   658 \item xxx
   659 \item xxx
   660 \end{itemize}
   661 \<close>
   662 
   663 section \<open>For copy & paste\<close>
   664 text \<open>
   665 \begin{itemize}
   666 \item xxx
   667   \begin{itemize}
   668   \item xxx
   669     \begin{itemize}
   670     \item xxx
   671       \begin{itemize}
   672       \item xxx
   673         \begin{itemize}
   674         \item xxx
   675         \end{itemize}
   676       \end{itemize}
   677     \end{itemize}
   678   \end{itemize}
   679 \end{itemize}
   680 \<close>
   681 subsection \<open>xxx\<close>
   682 subsubsection \<open>xxx\<close>
   683 end