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