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