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