src/Tools/isac/Build_Isac.thy
author Walther Neuper <walther.neuper@jku.at>
Mon, 11 May 2020 12:25:52 +0200
changeset 59963 e3cf90168a49
parent 59953 933211a252f2
child 59997 46fe5a8c3911
permissions -rw-r--r--
introduce Pre_Conds.T
     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           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
   138 
   139 (*//-----------------------------------------------------------------------------------------\\*)
   140 (*\\-----------------------------------------------------------------------------------------//*)
   141 begin
   142 
   143 text \<open>
   144   show theory dependencies using the graph browser, 
   145   open "browser_info/HOL/Isac/session.graph"
   146   and proceed from the ancestors towards the siblings.
   147 \<close>
   148 
   149 section \<open>check presence of definitions from directories\<close>
   150 
   151 (*declare [[ML_print_depth = 999]]*)
   152 ML \<open>
   153 \<close> ML \<open>
   154 \<close> ML \<open>
   155 \<close>
   156 ML \<open>Eval.adhoc_thm; (*from "ProgLang/evaluate.sml" *)\<close>
   157 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
   158 ML \<open>Input_Descript.for_real_list; (*from "Input_Descript.thy" *)\<close>
   159 ML \<open>Test_Code.me;\<close>
   160 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
   161 ML \<open>prog_expr\<close>
   162                                                          
   163 ML \<open>Prog_Expr.eval_occurs_in\<close>
   164 ML \<open>@{thm last_thmI}\<close>
   165 ML \<open>@{thm Querkraft_Belastung}\<close>
   166 
   167 ML \<open>Check_Unique.on := false;\<close>
   168 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
   169 ML \<open>@{theory "Isac_Knowledge"}\<close>
   170 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   171   ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*)\<close>
   172 
   173 section \<open>State of approaching Isabelle by Isac\<close>
   174 text \<open>
   175   Mathias Lehnfeld gives the following list in his thesis in section 
   176   4.2.3 Relation to Ongoing Isabelle Development.
   177 \<close>
   178 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   179 text \<open>
   180   REPLACE BY Know_Store... (has been overlooked)
   181     calcelems.sml:val rew_ord' = Unsynchronized.ref ...
   182   KEEP FOR TRACING
   183     rewrite.sml:val Rewrite.trace_on = Unsynchronized.ref false;
   184     rewrite.sml:val depth = Unsynchronized.ref 99999;
   185     rewrite.sml:val lim_rewrite = Unsynchronized.ref 99999;
   186     rewrite.sml:val lim_deriv = Unsynchronized.ref 100;
   187     Interpret/rewtools.sml:val LItool.trace = Unsynchronized.ref false;
   188   KEEP FOR EASIER DEVELOPMENT
   189     calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
   190   KEEP FOR DEMOS
   191     Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
   192     Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
   193     Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
   194 \<close>
   195 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   196 subsection \<open>(2) Make Isac’s programming language usable\<close>
   197 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   198 text \<open>
   199   In 2002 isac already strived for floating point numbers. Since that time
   200   isac represents numerals as "Free", see below (*1*). These experiments are
   201   unsatisfactory with respect to logical soundness.
   202   Since Isabelle now has started to care about floating point numbers, it is high 
   203   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   204   are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
   205   
   206   The transition from "Free" to standard numerals is a task to be scheduled for 
   207   several weeks. The urgency of this task follows from the experience, 
   208   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   209   some of the long identifiers of theorems which would tremendously simplify
   210   building a hierarchy of theorems according to (1.2), see (*2*) below.
   211 \<close>
   212 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   213 
   214 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   215 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   216 
   217 end