src/Tools/isac/Build_Isac.thy
author Walther Neuper <walther.neuper@jku.at>
Fri, 10 Apr 2020 14:46:55 +0200
changeset 59865 75a9d629ea53
parent 59858 a2c32a38327a
child 59866 3b194392ea71
permissions -rw-r--r--
rearrange code for ThmC
     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 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 CalcElements imports KEStore
    29     ML_file termC.sml
    30     ML_file contextC.sml
    31     ML_file environment.sml
    32 *)        "CalcElements/CalcElements"
    33 
    34 (*    theory Calculate imports "~~/src/Tools/isac/CalcElements/CalcElements"
    35         ML_file calculate.sml
    36 
    37       theory ListC imports "~~/src/Tools/isac/CalcElements/CalcElements"
    38     theory Prog_Expr imports Calculate ListC
    39       theory Program imports "~~/src/Tools/isac/CalcElements/CalcElements"
    40       theory Prog_Tac imports "~~/src/Tools/isac/CalcElements/CalcElements"
    41       theory Tactical imports "~~/src/Tools/isac/CalcElements/CalcElements"
    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 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/CalcElements/CalcElements"
    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 parseNEW @{context} "xxx #> aaa bbb"
   124 \<close> ML \<open>
   125 #> : ('a -> 'b) * ('b -> 'c) -> 'a -> 'c
   126 \<close> ML \<open>
   127 \<close>
   128 ML \<open>Num_Calc.adhoc_thm; (*from "ProgLang/calculate.sml" *)\<close>
   129 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
   130 ML \<open>Input_Descript.is_reall_dsc; (*from "ProgLang/scrtools.sml" *)\<close>
   131 ML \<open>Test_Code.me;\<close>
   132 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
   133 ML \<open>prog_expr\<close>
   134                                                          
   135 ML \<open>Prog_Expr.eval_occurs_in\<close>
   136 ML \<open>@{thm last_thmI}\<close>
   137 ML \<open>@{thm Querkraft_Belastung}\<close>
   138 
   139 ML \<open>Celem.check_guhs_unique := false;\<close>
   140 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
   141 ML \<open>@{theory "Isac_Knowledge"}\<close>
   142 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   143   ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*)\<close>
   144 
   145 section \<open>State of approaching Isabelle by Isac\<close>
   146 text \<open>
   147   Mathias Lehnfeld gives the following list in his thesis in section 
   148   4.2.3 Relation to Ongoing Isabelle Development.
   149 \<close>
   150 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   151 text \<open>
   152   REPLACE BY KEStore... (has been overlooked)
   153     calcelems.sml:val rew_ord' = Unsynchronized.ref ...
   154   KEEP FOR TRACING
   155     calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
   156     calcelems.sml:val depth = Unsynchronized.ref 99999;
   157     calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
   158     calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
   159     Interpret/script.sml:val trace_LI = Unsynchronized.ref false;
   160   KEEP FOR EASIER DEVELOPMENT
   161     calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
   162   KEEP FOR DEMOS
   163     Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
   164     Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
   165     Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
   166 \<close>
   167 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   168 subsection \<open>(2) Make Isac’s programming language usable\<close>
   169 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   170 text \<open>
   171   In 2002 isac already strived for floating point numbers. Since that time
   172   isac represents numerals as "Free", see below (*1*). These experiments are
   173   unsatisfactory with respect to logical soundness.
   174   Since Isabelle now has started to care about floating point numbers, it is high 
   175   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   176   are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
   177   
   178   The transition from "Free" to standard numerals is a task to be scheduled for 
   179   several weeks. The urgency of this task follows from the experience, 
   180   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   181   some of the long identifiers of theorems which would tremendously simplify
   182   building a hierarchy of theorems according to (1.2), see (*2*) below.
   183 \<close>
   184 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   185 
   186 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   187 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   188 
   189 end