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