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