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