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