src/Tools/isac/Build_Isac.thy
author Walther Neuper <walther.neuper@jku.at>
Tue, 10 Sep 2019 17:15:41 +0200
changeset 59614 3a6e953d2a36
parent 59613 8d28eab80f7f
child 59632 a0e0dc864fbd
permissions -rw-r--r--
Isabelle2018->19: unclarified "Exception- Size raised" in Build_Thydata.thy

note: this exception exists since several changesets;
but the exn is passed over, because "isabelle build" succeeds.
     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           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
    75 
    76 (*//-----------------------------------------------------------------------------------------\\*)
    77 (*\\-----------------------------------------------------------------------------------------//*)
    78 begin
    79 
    80 text \<open>
    81   show theory dependencies using the graph browser, 
    82   open "browser_info/HOL/Isac/session.graph"
    83   and proceed from the ancestors towards the siblings.
    84 \<close>
    85 
    86 section \<open>check presence of definitions from directories\<close>
    87 
    88 (*declare [[ML_print_depth = 999]]*)
    89 ML \<open>
    90 \<close> ML \<open>
    91 \<close> ML \<open>
    92 \<close>
    93 ML \<open>Calc.adhoc_thm; (*from "ProgLang/calculate.sml" *)\<close>
    94 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
    95 ML \<open>Input_Descript.is_reall_dsc; (*from "ProgLang/scrtools.sml" *)\<close>
    96 ML \<open>Math_Engine.me;\<close>
    97 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
    98 ML \<open>list_rls\<close>
    99                                                          
   100 ML \<open>Prog_Expr.eval_occurs_in\<close>
   101 ML \<open>@{thm last_thmI}\<close>
   102 ML \<open>@{thm Querkraft_Belastung}\<close>
   103 
   104 ML \<open>Celem.check_guhs_unique := false;\<close>
   105 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
   106 ML \<open>@{theory "Isac_Knowledge"}\<close>
   107 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   108   ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*)\<close>
   109 
   110 section \<open>State of approaching Isabelle by Isac\<close>
   111 text \<open>
   112   Mathias Lehnfeld gives the following list in his thesis in section 
   113   4.2.3 Relation to Ongoing Isabelle Development.
   114 \<close>
   115 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   116 text \<open>
   117   REPLACE BY KEStore... (has been overlooked)
   118     calcelems.sml:val rew_ord' = Unsynchronized.ref ...
   119   KEEP FOR TRACING
   120     calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
   121     calcelems.sml:val depth = Unsynchronized.ref 99999;
   122     calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
   123     calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
   124     Interpret/script.sml:val trace_script = Unsynchronized.ref false;
   125   KEEP FOR EASIER DEVELOPMENT
   126     calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
   127   KEEP FOR DEMOS
   128     Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
   129     Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
   130     Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
   131 \<close>
   132 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   133 subsection \<open>(2) Make Isac’s programming language usable\<close>
   134 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   135 text \<open>
   136   In 2002 isac already strived for floating point numbers. Since that time
   137   isac represents numerals as "Free", see below (*1*). These experiments are
   138   unsatisfactory with respect to logical soundness.
   139   Since Isabelle now has started to care about floating point numbers, it is high 
   140   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   141   are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
   142   
   143   The transition from "Free" to standard numerals is a task to be scheduled for 
   144   several weeks. The urgency of this task follows from the experience, 
   145   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   146   some of the long identifiers of theorems which would tremendously simplify
   147   building a hierarchy of theorems according to (1.2), see (*2*) below.
   148 \<close>
   149 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   150 
   151 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   152 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   153 
   154 end