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