src/Tools/isac/Build_Isac.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Thu, 15 Mar 2018 12:42:04 +0100
changeset 59406 509d70b507e5
parent 59389 627d25067f2f
child 59456 d56b817fbb14
permissions -rw-r--r--
separate structure Celem: CALC_ELEMENT, finished on src/
     1 (*  Title:  build the Isac Mathengine
     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 (theory dependencies!)
     7 
     8 ATTENTION: no errors in this theory do not mean that there are no errors in Isac;
     9 errors are rigorously detected when creating a heap.
    10 *)
    11 
    12 theory Build_Isac 
    13 imports
    14 (* structure inherited from migration which began with Isabelle2009. improve?
    15           theory KEStore
    16             ML_file "~~/src/Tools/isac/library.sml"
    17             ML_file "~~/src/Tools/isac/calcelems.sml"
    18         theory ListC 
    19           imports "~~/src/Tools/isac/KEStore"
    20           ML_file "~~/src/Tools/isac/ProgLang/termC.sml"
    21           ML_file "~~/src/Tools/isac/ProgLang/calculate.sml"
    22           ML_file "~~/src/Tools/isac/ProgLang/rewrite.sml"
    23       theory Tools imports ListC begin
    24     theory Script imports Tools begin
    25   theory ProgLang imports Script
    26     ML_file "~~/src/Tools/isac/ProgLang/scrtools.sml"
    27 *)        "ProgLang/ProgLang"
    28 
    29 (* use    "Interpret/mstools.sml"
    30    use    "Interpret/ctree.sml"
    31    use    "Interpret/ptyps.sml"
    32    use    "Interpret/generate.sml"                               
    33    use    "Interpret/calchead.sml"
    34    use    "Interpret/appl.sml"
    35    use    "Interpret/rewtools.sml"
    36    use    "Interpret/script.sml"
    37    use    "Interpret/solve.sml"
    38    use    "Interpret/inform.sml"
    39    use    "Interpret/mathengine.sml"
    40 *)        "Interpret/Interpret"
    41 
    42 (* use    "xmlsrc/mathml.sml"
    43    use    "xmlsrc/datatypes.sml"
    44    use    "xmlsrc/pbl-met-hierarchy.sml"
    45    use    "xmlsrc/thy-hierarchy.sml"                               
    46    use    "xmlsrc/interface-xml.sml"
    47 *)        "xmlsrc/xmlsrc"
    48 
    49 (* use     "Frontend/messages.sml"
    50    use     "Frontend/states.sml"
    51    use     "Frontend/interface.sml"
    52 
    53    use     "print_exn_G.sml"
    54 *)         "Frontend/Frontend"
    55 
    56            "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
    57 
    58            (* the Protocol for the connection isac-java -- Isabelle/Isac is built separately:
    59               libisabelle$ /usr/local/isabisac/bin/isabelle build -D . -bv
    60               here we partly check consistency libisabelle/Isac -- Isabelle/Isac: *)
    61            (*"~~/libisabelle-protocol/isabelle-2015/Protocol"*)
    62 begin
    63 
    64 text {* 
    65   show theory dependencies using the graph browser, 
    66   open "browser_info/HOL/Isac/session.graph"
    67   and proceed from the ancestors towards the siblings.
    68 *}
    69 
    70 section {*check presence of definitions from directories*}
    71 
    72 (*declare [[ML_print_depth = 999]]*)
    73 ML {*
    74 *} ML {*
    75 *}
    76 
    77 (*==============================================================================================
    78   The test below from calculate.sml raises an exception with the cleaned Rewrite.
    79   The differences to the working 'fun rewrite' are that significant,
    80   that we want to rely on 'hg revert'.
    81   For that purpose we save this changeset with a broken 'fun rewrite' and tests not running.
    82 *)
    83 ML {*
    84 *} ML {*
    85 (*--------------(2): does divide work in Test_simplify ?: ------*)
    86  val thy = @{theory Test};
    87  val t = (Thm.term_of o the o (TermC.parse thy)) "6 / 2";
    88  val rls = Test_simplify;
    89 *} ML {*
    90  val (t,_) = the (Rewrite.rewrite_set_ thy false rls t);
    91 (*val t = Free ("3","Real.real") : term*)
    92 *} ML {*
    93 
    94 (*--------------(3): is_const works ?: -------------------------------------*)
    95  val t = (Thm.term_of o the o (TermC.parse @{theory})) "2 is_const";
    96 *} ML {*
    97  Rewrite.rewrite_set_ @{theory Test} false tval_rls t;
    98 
    99 (* exception TERM raised (line 270 of "~~/src/HOL/Tools/hologic.ML"):
   100   dest_eq
   101   0 \<le> ?a \<Longrightarrow> (0 \<le> sqrt ?a) = True*)
   102 *} ML {*
   103 *}
   104 (*==============================================================================================*)
   105 
   106 ML {* Calc.adhoc_thm; (*from "ProgLang/calculate.sml" *) *}
   107 ML {* Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *) *}
   108 ML {* LTool.is_reall_dsc; (*from "ProgLang/scrtools.sml" *) *}
   109 ML {* Math_Engine.me; (*from "Interpret/mathengine.sml"*) *}
   110 ML {* contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*) *}
   111 ML {* print_exn_unit *}
   112 ML {* list_rls (*from Atools.thy WN130615??? or ProgLang???*) *}
   113 
   114 ML {* eval_occurs_in (*from Atools.thy*) *}
   115 ML {* @{thm last_thmI} (*from Atools.thy*) *}
   116 ML {*@{thm Querkraft_Belastung}*}
   117 
   118 ML {* Celem.check_guhs_unique := false; *}
   119 ML {* writeln "**** isac kernel = math-engine + Knowledge complete ******" *}
   120 ML {* @{theory "Isac"} *}
   121 ML {* (*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   122   ERROR: app_py: not found: ["IsacKnowledge","Diff","Theorems","diff_sin_chain"]*) *}
   123 
   124 section {* State of approaching Isabelle by Isac *}
   125 text {* 
   126   Mathias Lehnfeld gives the following list in his thesis in section 
   127   4.2.3 Relation to Ongoing Isabelle Development.
   128 *}
   129 subsection {* (0) Survey on remaining Unsynchronized.ref *}
   130 text {*
   131   REPLACE BY KEStore... (has been overlooked)
   132     calcelems.sml:val rew_ord' = Unsynchronized.ref ...
   133   KEEP FOR TRACING
   134     calcelems.sml:val trace_rewrite = Unsynchronized.ref false;
   135     calcelems.sml:val depth = Unsynchronized.ref 99999;
   136     calcelems.sml:val lim_rewrite = Unsynchronized.ref 99999;
   137     calcelems.sml:val lim_deriv = Unsynchronized.ref 100;
   138     Interpret/script.sml:val trace_script = Unsynchronized.ref false;
   139   KEEP FOR EASIER DEVELOPMENT
   140     calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
   141   KEEP FOR DEMOS
   142     Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
   143     Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
   144     Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
   145 *}
   146 subsection {* (1) Exploit parallelism for concurrent session management *}
   147 subsection {* (2) Make Isac’s programming language usable *}
   148 subsection {* (3) Adopt Isabelle’s numeral computation for Isac *}
   149 text {* 
   150   In 2002 isac already strived for floating point numbers. Since that time
   151   isac represents numerals as "Free", see below (*1*). These experiments are
   152   unsatisfactory with respect to logical soundness.
   153   Since Isabelle now has started to care about floating point numbers, it is high 
   154   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   155   are different from Isabelle2011, see "test/Tools/isac/ProgLang/termC.sml".
   156   
   157   The transition from "Free" to standard numerals is a task to be scheduled for 
   158   several weeks. The urgency of this task follows from the experience, 
   159   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   160   some of the long identifiers of theorems which would tremendously simplify
   161   building a hierarchy of theorems according to (1.2), see (*2*) below.
   162 *}
   163 ML {*(*1*) Free ("123.456", HOLogic.realT) *}
   164 
   165 subsection {* (4) Improve the efficiency of Isac’s rewrite-engine *}
   166 subsection {* (5) Adopt Isabelle/jEdit for Isac *}
   167 
   168 end