src/Tools/isac/Build_Isac.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Thu, 11 Jul 2019 16:39:46 +0200
changeset 59568 c950cc4db043
parent 59562 d50fe358f04a
child 59577 60d191402598
permissions -rw-r--r--
add update for istate, ctxt to Ctree

note: related functions need revision;
these do not reflect an elder design change/
     1 (*  Title:  build the Isac Mathengine
     2     Author: Walther Neuper, TU Graz, 100808
     3    (c) due to copyright terms
     4 
     5 For creating a heap image of isac see ~~/ROOT.
     6 For debugging see text at begin (theory dependencies!)
     7 
     8 ATTENTION: no errors in this theory do not mean that there are no errors in Isac;
     9 errors are rigorously detected when creating a heap.
    10 *)
    11 
    12 theory Build_Isac 
    13 imports
    14 (* structure inherited from migration which began with Isabelle2009. improve?
    15             theory KEStore
    16               ML_file "~~/src/Tools/isac/library.sml"
    17               ML_file "~~/src/Tools/isac/ThydataC/rule.sml"
    18               ML_file "~~/src/Tools/isac/calcelems.sml"
    19           theory ListC 
    20             imports "~~/src/Tools/isac/KEStore"
    21             ML_file "~~/src/Tools/isac/ProgLang/termC.sml"
    22             ML_file "~~/src/Tools/isac/ProgLang/calculate.sml"
    23             ML_file "~~/src/Tools/isac/ProgLang/rewrite.sml"
    24         theory Tools imports ListC begin
    25       theory Script imports Tools begin
    26     theory Atools imports Descript Script
    27       ML_file "~~/src/Tools/isac/ProgLang/scrtools.sml"
    28   theory ProgLang imports Atools
    29 *)        "ProgLang/ProgLang"
    30 
    31 (*ML_file "~~/src/Tools/isac/Interpret/model.sml"
    32   ML_file "~~/src/Tools/isac/Interpret/mstools.sml"
    33   ML_file "~~/src/Tools/isac/Interpret/specification-elems.sml"
    34   ML_file "~~/src/Tools/isac/Interpret/tactic.sml"
    35   ML_file "~~/src/Tools/isac/Interpret/ctree-basic.sml"
    36   ML_file "~~/src/Tools/isac/Interpret/ctree-access.sml"
    37   ML_file "~~/src/Tools/isac/Interpret/ctree-navi.sml"
    38   ML_file "~~/src/Tools/isac/Interpret/ctree.sml"
    39   ML_file "~~/src/Tools/isac/Interpret/ptyps.sml"
    40   ML_file "~~/src/Tools/isac/Interpret/generate.sml"
    41   ML_file "~~/src/Tools/isac/Interpret/calchead.sml"
    42   ML_file "~~/src/Tools/isac/Interpret/appl.sml"
    43   ML_file "~~/src/Tools/isac/Interpret/rewtools.sml"
    44   ML_file "~~/src/Tools/isac/Interpret/script.sml"
    45   ML_file "~~/src/Tools/isac/Interpret/inform.sml"
    46   ML_file "~~/src/Tools/isac/Interpret/lucas-interpreter.sml"
    47   ML_file "~~/src/Tools/isac/Interpret/solve.sml"
    48   ML_file "~~/src/Tools/isac/Interpret/mathengine.sml"
    49 *)        "Interpret/Interpret"
    50 
    51 (*ML_file "~~/src/Tools/isac/xmlsrc/mathml.sml"
    52   ML_file "~~/src/Tools/isac/xmlsrc/datatypes.sml"
    53   ML_file "~~/src/Tools/isac/xmlsrc/pbl-met-hierarchy.sml"
    54   ML_file "~~/src/Tools/isac/xmlsrc/thy-hierarchy.sml" 
    55   ML_file "~~/src/Tools/isac/xmlsrc/interface-xml.sml"
    56 *)        "xmlsrc/xmlsrc"
    57 
    58 (*ML_file "~~/src/Tools/isac/Frontend/messages.sml"
    59   ML_file "~~/src/Tools/isac/Frontend/states.sml"
    60   ML_file "~~/src/Tools/isac/Frontend/interface.sml"
    61 
    62   ML_file "~~/src/Tools/isac/print_exn_G.sml"
    63 *)         "Frontend/Frontend"
    64 
    65            "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
    66 
    67            (* the Protocol for the connection isac-java -- Isabelle/Isac is built separately:
    68               here we partly check consistency libisabelle/Isac -- Isabelle/Isac: *)
    69            Protocol.Protocol
    70 
    71 begin
    72 
    73 text \<open>
    74   show theory dependencies using the graph browser, 
    75   open "browser_info/HOL/Isac/session.graph"
    76   and proceed from the ancestors towards the siblings.
    77 \<close>
    78 
    79 section \<open>check presence of definitions from directories\<close>
    80 
    81 (*declare [[ML_print_depth = 999]]*)
    82 ML \<open>
    83 \<close> ML \<open>
    84 \<close>
    85 ML \<open>Calc.adhoc_thm; (*from "ProgLang/calculate.sml" *)\<close>
    86 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
    87 ML \<open>LTool.is_reall_dsc; (*from "ProgLang/scrtools.sml" *)\<close>
    88 ML \<open>Math_Engine.me; (*from "Interpret/mathengine.sml"*)\<close>
    89 ML \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
    90 ML \<open>print_exn_unit\<close>
    91 ML \<open>list_rls (*from Atools.thy WN130615??? or ProgLang???*)\<close>
    92                                                          
    93 ML \<open>eval_occurs_in (*from Atools.thy*)\<close>
    94 ML \<open>@{thm last_thmI} (*from Atools.thy*)\<close>
    95 ML \<open>@{thm Querkraft_Belastung}\<close>
    96 
    97 ML \<open>Celem.check_guhs_unique := false;\<close>
    98 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
    99 ML \<open>@{theory "Isac"}\<close>
   100 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   101   ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*)\<close>
   102 
   103 section \<open>State of approaching Isabelle by Isac\<close>
   104 text \<open>
   105   Mathias Lehnfeld gives the following list in his thesis in section 
   106   4.2.3 Relation to Ongoing Isabelle Development.
   107 \<close>
   108 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   109 text \<open>
   110   REPLACE BY KEStore... (has been overlooked)
   111     calcelems.sml:val rew_ord' = Unsynchronized.ref ...
   112   KEEP FOR TRACING
   113     calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
   114     calcelems.sml:val depth = Unsynchronized.ref 99999;
   115     calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
   116     calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
   117     Interpret/script.sml:val trace_script = Unsynchronized.ref false;
   118   KEEP FOR EASIER DEVELOPMENT
   119     calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
   120   KEEP FOR DEMOS
   121     Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
   122     Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
   123     Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
   124 \<close>
   125 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   126 subsection \<open>(2) Make Isac’s programming language usable\<close>
   127 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   128 text \<open>
   129   In 2002 isac already strived for floating point numbers. Since that time
   130   isac represents numerals as "Free", see below (*1*). These experiments are
   131   unsatisfactory with respect to logical soundness.
   132   Since Isabelle now has started to care about floating point numbers, it is high 
   133   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   134   are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
   135   
   136   The transition from "Free" to standard numerals is a task to be scheduled for 
   137   several weeks. The urgency of this task follows from the experience, 
   138   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   139   some of the long identifiers of theorems which would tremendously simplify
   140   building a hierarchy of theorems according to (1.2), see (*2*) below.
   141 \<close>
   142 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   143 
   144 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   145 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   146 
   147 end