src/Tools/isac/Build_Isac.thy
author Walther Neuper <walther.neuper@jku.at>
Tue, 10 Sep 2019 16:13:28 +0200
changeset 59613 8d28eab80f7f
parent 59612 14b7eae04d42
child 59614 3a6e953d2a36
permissions -rw-r--r--
Isabelle2018->19: rm libisabelle finished, retain interface.sml

note: outcommented "Exception- Size raised" in Build_Thydata.thy
     1 (*  Title:  build the Isac 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 (theory dependencies!)
     7 
     8 ATTENTION: no errors in this theory do not mean that there are no errors in Isac ..
     9 .. open theories collecting files from folders: CalcElements.thy, ProgLang.thy etc.
    10 Errors are rigorously detected when creating a heap.
    11 *)
    12 
    13 theory Build_Isac 
    14 imports
    15 (* see dependency graph
    16   ~$ evince file:///home/wneuper/.isabelle/isabisac/browser_info/Unsorted/Isac/session_graph.pdf &
    17 *)
    18 (*    theory KEStore imports Complex_Main
    19         ML_file "~~/src/Tools/isac/CalcElements/libraryC.sml"
    20         ML_file "~~/src/Tools/isac/CalcElements/rule.sml"
    21         ML_file "~~/src/Tools/isac/CalcElements/calcelems.sml"
    22     theory ListC imports KEStore
    23       ML_file "~~/src/Tools/isac/CalcElements/termC.sml"
    24       ML_file "~~/src/Tools/isac/CalcElements/contextC.sml"
    25   theory CalcElements imports ListC
    26 *)        "CalcElements/CalcElements"
    27 
    28 (*      theory Tools imports CalcElements begin
    29           ML_file "~~/src/Tools/isac/ProgLang/calculate.sml"
    30           ML_file "~~/src/Tools/isac/ProgLang/rewrite.sml"
    31       theory Program imports Tools begin
    32         ML_file "~~/src/Tools/isac/ProgLang/program.sml" ? really separate?
    33     theory Atools imports Program
    34       ML_file "~~/src/Tools/isac/ProgLang/scrtools.sml"
    35   theory ProgLang imports Atools
    36 *)        "ProgLang/ProgLang"
    37 (*
    38   ML_file model.sml
    39   ML_file mstools.sml
    40   ML_file "specification-elems.sml"
    41   ML_file istate.sml
    42   ML_file tactic.sml
    43   ML_file "ctree-basic.sml" (*shift to base in common with Interpret*)
    44   ML_file "ctree-access.sml"(*shift to base in common with Interpret*)
    45   ML_file "ctree-navi.sml"  (*shift to base in common with Interpret*)
    46   ML_file ctree.sml         (*shift to base in common with Interpret*)
    47   ML_file ptyps.sml
    48   ML_file generate.sml
    49   ML_file calchead.sml
    50   ML_file appl.sml
    51 *)        "Specify/Specify"
    52 (*
    53   ML_file rewtools.sml
    54   ML_file script.sml
    55   ML_file inform.sml
    56   ML_file "lucas-interpreter.sml"
    57 *)        "Interpret/Interpret"
    58 (*
    59   ML_file solve.sml
    60   ML_file mathengine.sml
    61   ML_file "~~/src/Tools/isac/Frontend/messages.sml"
    62   ML_file messages.sml
    63   ML_file states.sml
    64 *)        "MathEngine/MathEngine"
    65 (*
    66   ML_file mathml.sml
    67   ML_file datatypes.sml
    68   ML_file "pbl-met-hierarchy.sml"
    69   ML_file "thy-hierarchy.sml" 
    70   ML_file "interface-xml.sml"
    71   ML_file interface.sml
    72           "BridgeLibisabelle/BridgeLibisabelle"
    73 *)
    74 
    75           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
    76 
    77 (*//-----------------------------------------------------------------------------------------\\*)
    78 (*\\-----------------------------------------------------------------------------------------//*)
    79 begin
    80 
    81 text \<open>
    82   show theory dependencies using the graph browser, 
    83   open "browser_info/HOL/Isac/session.graph"
    84   and proceed from the ancestors towards the siblings.
    85 \<close>
    86 
    87 section \<open>check presence of definitions from directories\<close>
    88 
    89 (*declare [[ML_print_depth = 999]]*)
    90 ML \<open>
    91 \<close> ML \<open>
    92 \<close> ML \<open>
    93 \<close>
    94 ML \<open>Calc.adhoc_thm; (*from "ProgLang/calculate.sml" *)\<close>
    95 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
    96 ML \<open>Input_Descript.is_reall_dsc; (*from "ProgLang/scrtools.sml" *)\<close>
    97 ML \<open>Math_Engine.me;\<close>
    98 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
    99 ML \<open>list_rls\<close>
   100                                                          
   101 ML \<open>Prog_Expr.eval_occurs_in\<close>
   102 ML \<open>@{thm last_thmI}\<close>
   103 ML \<open>@{thm Querkraft_Belastung}\<close>
   104 
   105 ML \<open>Celem.check_guhs_unique := false;\<close>
   106 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
   107 ML \<open>@{theory "Isac_Knowledge"}\<close>
   108 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   109   ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*)\<close>
   110 
   111 section \<open>State of approaching Isabelle by Isac\<close>
   112 text \<open>
   113   Mathias Lehnfeld gives the following list in his thesis in section 
   114   4.2.3 Relation to Ongoing Isabelle Development.
   115 \<close>
   116 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   117 text \<open>
   118   REPLACE BY KEStore... (has been overlooked)
   119     calcelems.sml:val rew_ord' = Unsynchronized.ref ...
   120   KEEP FOR TRACING
   121     calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
   122     calcelems.sml:val depth = Unsynchronized.ref 99999;
   123     calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
   124     calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
   125     Interpret/script.sml:val trace_script = Unsynchronized.ref false;
   126   KEEP FOR EASIER DEVELOPMENT
   127     calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
   128   KEEP FOR DEMOS
   129     Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
   130     Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
   131     Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
   132 \<close>
   133 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   134 subsection \<open>(2) Make Isac’s programming language usable\<close>
   135 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   136 text \<open>
   137   In 2002 isac already strived for floating point numbers. Since that time
   138   isac represents numerals as "Free", see below (*1*). These experiments are
   139   unsatisfactory with respect to logical soundness.
   140   Since Isabelle now has started to care about floating point numbers, it is high 
   141   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   142   are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
   143   
   144   The transition from "Free" to standard numerals is a task to be scheduled for 
   145   several weeks. The urgency of this task follows from the experience, 
   146   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   147   some of the long identifiers of theorems which would tremendously simplify
   148   building a hierarchy of theorems according to (1.2), see (*2*) below.
   149 \<close>
   150 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   151 
   152 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   153 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   154 
   155 end