src/Tools/isac/Build_Isac.thy
author Walther Neuper <walther.neuper@jku.at>
Sun, 22 Sep 2019 15:15:37 +0200
changeset 59632 a0e0dc864fbd
parent 59614 3a6e953d2a36
child 59633 f854e130f851
permissions -rw-r--r--
lucin: update comments on structure of import's and ML_file's
     1 (*  Title:  build the Isac-Kernel: MathEngine & Knowledge
     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 below, e.g. theory dependencies:
     7   ~$ evince file:///home/wneuper/.isabelle/isabisac/browser_info/Unsorted/Isac/session_graph.pdf &
     8 
     9 ATTENTION: no errors in this theory do not mean that there are no errors in Isac ..
    10 .. open theories collecting files from folders: CalcElements.thy, ProgLang.thy etc.
    11 Errors are rigorously detected by isabelle build.
    12 *)
    13 
    14 theory Build_Isac 
    15 imports
    16 (*  theory KEStore imports Complex_Main
    17       ML_file libraryC.sml
    18       ML_file rule.sml
    19       ML_file calcelems.sml
    20   theory CalcElements imports KEStore
    21     ML_file termC.sml
    22     ML_file contextC.sml
    23 *)        "CalcElements/CalcElements"
    24 
    25 (*    theory Calculate imports "~~/src/Tools/isac/CalcElements/CalcElements"
    26         ML_file calculate.sml
    27       theory ListC imports "~~/src/Tools/isac/CalcElements/CalcElements"
    28     theory Prog_Expr imports Calculate ListC
    29 
    30       theory Program imports "~~/src/Tools/isac/CalcElements/CalcElements"
    31       theory Prog_Tac imports "~~/src/Tools/isac/CalcElements/CalcElements"
    32       theory Tactical imports "~~/src/Tools/isac/CalcElements/CalcElements"
    33     theory Auto_Prog imports Program Prog_Tac Tactical begin
    34   theory ProgLang imports Prog_Expr Auto_Prog
    35     ML_file rewrite.sml
    36 *)        "ProgLang/ProgLang"
    37 (*
    38     theory Input_Descript imports "~~/src/Tools/isac/CalcElements/CalcElements"
    39   theory Specify imports "~~/src/Tools/isac/ProgLang/ProgLang" Input_Descript
    40     ML_file model.sml
    41     ML_file mstools.sml
    42     ML_file "specification-elems.sml"
    43     ML_file istate.sml
    44     ML_file tactic.sml
    45     ML_file "ctree-basic.sml" (*shift to base in common with Interpret*)
    46     ML_file "ctree-access.sml"(*shift to base in common with Interpret*)
    47     ML_file "ctree-navi.sml"  (*shift to base in common with Interpret*)
    48     ML_file ctree.sml         (*shift to base in common with Interpret*)
    49     ML_file ptyps.sml
    50     ML_file generate.sml
    51     ML_file calchead.sml
    52     ML_file appl.sml
    53 *)        "Specify/Specify"
    54 (*
    55   theory Interpret imports "~~/src/Tools/isac/Specify/Specify"
    56     ML_file rewtools.sml
    57     ML_file script.sml
    58     ML_file inform.sml
    59     ML_file "lucas-interpreter.sml"
    60 *)        "Interpret/Interpret"
    61 (*
    62   theory MathEngine imports "~~/src/Tools/isac/Interpret/Interpret"
    63     ML_file solve.sml
    64     ML_file mathengine.sml
    65     ML_file "~~/src/Tools/isac/Frontend/messages.sml"
    66     ML_file messages.sml
    67     ML_file states.sml
    68 *)        "MathEngine/MathEngine"
    69 (*
    70   theory BridgeLibisabelle imports "~~/src/Tools/isac/MathEngine/MathEngine"
    71     ML_file mathml.sml
    72     ML_file datatypes.sml
    73     ML_file "pbl-met-hierarchy.sml"
    74     ML_file "thy-hierarchy.sml" 
    75     ML_file "interface-xml.sml"
    76     ML_file interface.sml
    77 *)        "BridgeLibisabelle/BridgeLibisabelle"
    78 
    79           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
    80 
    81 (*//-----------------------------------------------------------------------------------------\\*)
    82 (*\\-----------------------------------------------------------------------------------------//*)
    83 begin
    84 
    85 text \<open>
    86   show theory dependencies using the graph browser, 
    87   open "browser_info/HOL/Isac/session.graph"
    88   and proceed from the ancestors towards the siblings.
    89 \<close>
    90 
    91 section \<open>check presence of definitions from directories\<close>
    92 
    93 (*declare [[ML_print_depth = 999]]*)
    94 ML \<open>
    95 \<close> ML \<open>
    96 \<close> ML \<open>
    97 \<close>
    98 ML \<open>Calc.adhoc_thm; (*from "ProgLang/calculate.sml" *)\<close>
    99 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
   100 ML \<open>Input_Descript.is_reall_dsc; (*from "ProgLang/scrtools.sml" *)\<close>
   101 ML \<open>Math_Engine.me;\<close>
   102 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
   103 ML \<open>list_rls\<close>
   104                                                          
   105 ML \<open>Prog_Expr.eval_occurs_in\<close>
   106 ML \<open>@{thm last_thmI}\<close>
   107 ML \<open>@{thm Querkraft_Belastung}\<close>
   108 
   109 ML \<open>Celem.check_guhs_unique := false;\<close>
   110 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
   111 ML \<open>@{theory "Isac_Knowledge"}\<close>
   112 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   113   ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*)\<close>
   114 
   115 section \<open>State of approaching Isabelle by Isac\<close>
   116 text \<open>
   117   Mathias Lehnfeld gives the following list in his thesis in section 
   118   4.2.3 Relation to Ongoing Isabelle Development.
   119 \<close>
   120 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   121 text \<open>
   122   REPLACE BY KEStore... (has been overlooked)
   123     calcelems.sml:val rew_ord' = Unsynchronized.ref ...
   124   KEEP FOR TRACING
   125     calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
   126     calcelems.sml:val depth = Unsynchronized.ref 99999;
   127     calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
   128     calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
   129     Interpret/script.sml:val trace_script = Unsynchronized.ref false;
   130   KEEP FOR EASIER DEVELOPMENT
   131     calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
   132   KEEP FOR DEMOS
   133     Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
   134     Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
   135     Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
   136 \<close>
   137 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   138 subsection \<open>(2) Make Isac’s programming language usable\<close>
   139 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   140 text \<open>
   141   In 2002 isac already strived for floating point numbers. Since that time
   142   isac represents numerals as "Free", see below (*1*). These experiments are
   143   unsatisfactory with respect to logical soundness.
   144   Since Isabelle now has started to care about floating point numbers, it is high 
   145   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   146   are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
   147   
   148   The transition from "Free" to standard numerals is a task to be scheduled for 
   149   several weeks. The urgency of this task follows from the experience, 
   150   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   151   some of the long identifiers of theorems which would tremendously simplify
   152   building a hierarchy of theorems according to (1.2), see (*2*) below.
   153 \<close>
   154 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   155 
   156 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   157 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   158 
   159 end