src/Tools/isac/Build_Isac.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Sun, 16 Oct 2016 13:58:46 +0200
changeset 59251 241e06eb9c04
parent 59206 ebf4a8a63371
child 59259 d1c13ee4e1a2
permissions -rw-r--r--
=== note on "tricky combination of (string, term) for theorems in Isac"

Notes:
# Test_Isac does not work due to Rewrite'..thm
# TODO undo Rewrite'..thm ---> Rewrite'..(thmID,_)
# approach revision of "fun assoc_thm'" stepwise (start with cleanup)
     1 (*  Title:  build and test isac on Isabelle2014
     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 
     7 ATTENTION: no errors in this theory do not mean that there are no errors in Isac;
     8 errors are rigorously detected when creating a heap.
     9 *)
    10 
    11 header {* Loading the isac mathengine;
    12   for debugging see text at begin (theory dependencies!) *}
    13 
    14 theory Build_Isac 
    15 imports
    16 (* structure inherited from migration which began with Isabelle2009. improve?
    17           theory KEStore
    18             ML_file "~~/src/Tools/isac/library.sml"
    19             ML_file "~~/src/Tools/isac/calcelems.sml"
    20         theory ListC 
    21           imports "~~/src/Tools/isac/KEStore"
    22           ML_file "~~/src/Tools/isac/ProgLang/termC.sml"
    23           ML_file "~~/src/Tools/isac/ProgLang/calculate.sml"
    24           ML_file "~~/src/Tools/isac/ProgLang/rewrite.sml"
    25       theory Tools imports ListC begin
    26     theory Script imports Tools begin
    27   theory ProgLang imports Script
    28     ML_file "~~/src/Tools/isac/ProgLang/scrtools.sml"
    29 *)        "ProgLang/ProgLang"
    30 
    31 (* use    "Interpret/mstools.sml"
    32    use    "Interpret/ctree.sml"
    33    use    "Interpret/ptyps.sml"
    34    use    "Interpret/generate.sml"                               
    35    use    "Interpret/calchead.sml"
    36    use    "Interpret/appl.sml"
    37    use    "Interpret/rewtools.sml"
    38    use    "Interpret/script.sml"
    39    use    "Interpret/solve.sml"
    40    use    "Interpret/inform.sml"
    41    use    "Interpret/mathengine.sml"
    42 *)        "Interpret/Interpret"
    43 
    44 (* use    "xmlsrc/mathml.sml"
    45    use    "xmlsrc/datatypes.sml"
    46    use    "xmlsrc/pbl-met-hierarchy.sml"
    47    use    "xmlsrc/thy-hierarchy.sml"                               
    48    use    "xmlsrc/interface-xml.sml"
    49 *)        "xmlsrc/xmlsrc"
    50 
    51 (* use     "Frontend/messages.sml"
    52    use     "Frontend/states.sml"
    53    use     "Frontend/interface.sml"
    54 
    55    use     "print_exn_G.sml"
    56 *)         "Frontend/Frontend"
    57 
    58            "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
    59 
    60            (* the Protocol for the connection isac-java -- Isabelle/Isac is built separately:
    61               libisabelle$ /usr/local/isabisac/bin/isabelle build -D . -bv
    62               here we partly check consistency libisabelle/Isac -- Isabelle/Isac: *)
    63            (*"~~/libisabelle-protocol/isabelle-2015/Protocol"*)
    64 begin
    65 
    66 text {* 
    67   show theory dependencies using the graph browser, 
    68   open "browser_info/HOL/Isac/session.graph"
    69   and proceed from the ancestors towards the siblings.
    70 *}
    71 
    72 section {*check presence of definitions from directories*}
    73 
    74 ML {*
    75 *} ML {*
    76 *}
    77 ML {* is_reall_dsc; (*from "ProgLang/scrtools.sml" *) *}
    78 ML {* me; (*from "Interpret/mathengine.sml"*) *}
    79 ML {* contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*) *}
    80 ML {* print_exn_unit *}
    81 ML {* list_rls (*from Atools.thy WN130615??? or ProgLang???*) *}
    82 
    83 ML {* eval_occurs_in (*from Atools.thy*) *}
    84 ML {* @{thm last_thmI} (*from Atools.thy*) *}
    85 ML {*@{thm Querkraft_Belastung}*}
    86 
    87 ML {* check_guhs_unique := false; *}
    88 ML {* writeln "**** isac kernel = math-engine + Knowledge complete ******" *}
    89 ML {* @{theory "Isac"} *}
    90 ML {* (*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
    91   ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*) *}
    92 
    93 section {* State of approaching Isabelle by Isac *}
    94 text {* 
    95   Mathias Lehnfeld gives the following list in his thesis in section 
    96   4.2.3 Relation to Ongoing Isabelle Development.
    97 *}
    98 subsection {* (0) Survey on remaining Unsynchronized.ref *}
    99 text {*
   100   REPLACE BY KEStore... (has been overlooked)
   101     calcelems.sml:val rew_ord' = Unsynchronized.ref ...
   102   KEEP FOR TRACING
   103     calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
   104     calcelems.sml:val depth = Unsynchronized.ref 99999;
   105     calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
   106     calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
   107     Interpret/script.sml:val trace_script = Unsynchronized.ref false;
   108   KEEP FOR EASIER DEVELOPMENT
   109     calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
   110   KEEP FOR DEMOS
   111     Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
   112     Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
   113     Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
   114 *}
   115 subsection {* (1) Exploit parallelism for concurrent session management *}
   116 subsection {* (2) Make Isac’s programming language usable *}
   117 subsection {* (3) Adopt Isabelle’s numeral computation for Isac *}
   118 text {* 
   119   In 2002 isac already strived for floating point numbers. Since that time
   120   isac represents numerals as "Free", see below (*1*). These experiments are
   121   unsatisfactory with respect to logical soundness.
   122   Since Isabelle now has started to care about floating point numbers, it is high 
   123   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   124   are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
   125   
   126   The transition from "Free" to standard numerals is a task to be scheduled for 
   127   several weeks. The urgency of this task follows from the experience, 
   128   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   129   some of the long identifiers of theorems which would tremendously simplify
   130   building a hierarchy of theorems according to (1.2), see (*2*) below.
   131 *}
   132 ML {*(*1*) Free ("123.456", HOLogic.realT) *}
   133 
   134 subsection {* (4) Improve the efficiency of Isac’s rewrite-engine *}
   135 subsection {* (5) Adopt Isabelle/jEdit for Isac *}
   136 
   137 end