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