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