src/Tools/isac/Build_Isac.thy
author Walther Neuper <walther.neuper@jku.at>
Wed, 29 Apr 2020 09:03:01 +0200
changeset 59919 3a7fb975af9d
parent 59918 58d9fcc5a712
child 59936 554030065b5b
permissions -rw-r--r--
comments on relation between files.

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