src/Tools/isac/Build_Isac.thy
author wenzelm
Sun, 18 Apr 2021 18:30:31 +0200
changeset 60217 1d9fee958a46
parent 60192 4c7c15750166
child 60317 638d02a9a96a
permissions -rw-r--r--
proper test sessions, but with remaining failures;
     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: BaseDefinitions.thy, ProgLang.thy etc.
    11 Errors are rigorously detected by isabelle build.
    12 *)
    13 
    14 theory Build_Isac
    15 imports
    16 (*  theory Know_Store imports Complex_Main
    17     $ISABELLE_ISAC/BaseDefinitions
    18       ML_file libraryC.sml
    19       ML_file theoryC.sml
    20       ML_file unparseC.sml
    21       ML_file "rule-def.sml"
    22       ML_file "thmC-def.sml"
    23       ML_file "eval-def.sml"
    24       ML_file "rewrite-order.sml"
    25       ML_file rule.sml
    26       ML_file "error-pattern-def.sml"
    27       ML_file "rule-set.sml"
    28       
    29       ML_file "store.sml"
    30       ML_file "check-unique.sml"
    31       ML_file "specification.sml"
    32       ML_file "model-pattern.sml"
    33       ML_file "problem-def.sml"
    34       ML_file "method-def.sml"
    35       ML_file "cas-def.sml"
    36       ML_file "thy-write.sml"
    37   theory BaseDefinitions imports Know_Store
    38   $ISABELLE_ISAC/BaseDefinitions
    39     ML_file termC.sml
    40     ML_file substitution.sml
    41     ML_file contextC.sml
    42     ML_file environment.sml
    43 ( ** )    "BaseDefinitions/BaseDefinitions"( **)
    44 
    45 (*    theory Calculate imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    46       $ISABELLE_ISAC/ProgLang
    47         ML_file evaluate.sml
    48 
    49       theory ListC imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    50       $ISABELLE_ISAC/ProgLang
    51     theory Prog_Expr imports Calculate ListC
    52     $ISABELLE_ISAC/ProgLang
    53       theory Program imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    54       theory Prog_Tac imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    55       theory Tactical imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    56     theory Auto_Prog imports Program Prog_Tac Tactical begin
    57     $ISABELLE_ISAC/ProgLang
    58   theory ProgLang imports Prog_Expr Auto_Prog
    59   $ISABELLE_ISAC/ProgLang
    60 ( ** )    "ProgLang/ProgLang"( **)
    61 (*
    62   theory MathEngBasic imports
    63     "$ISABELLE_ISAC/ProgLang/ProgLang" "$ISABELLE_ISAC/Specify/Input_Descript"
    64   $ISABELLE_ISAC/MathEngBasic
    65     ML_file thmC.sml
    66     ML_file problem.sml
    67     ML_file method.sml
    68     ML_file "cas-command.sml"
    69   
    70     ML_file rewrite.sml
    71   
    72     ML_file "model-def.sml"
    73     ML_file "istate-def.sml"
    74     ML_file "calc-tree-elem.sml"
    75     ML_file "pre-conds-def.sml"
    76   
    77     ML_file tactic.sml
    78     ML_file applicable.sml
    79   
    80     ML_file position.sml
    81     ML_file "ctree-basic.sml"
    82     ML_file "ctree-access.sml"
    83     ML_file "ctree-navi.sml"
    84     ML_file ctree.sml
    85   
    86     ML_file "state-steps.sml"
    87     ML_file calculation.sml
    88 ( ** )    "MathEngBasic/MathEngBasic"( **)
    89 (*
    90     theory Input_Descript imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    91     $ISABELLE_ISAC/Specify
    92   theory Specify imports "$ISABELLE_ISAC/ProgLang/ProgLang" Input_Descript
    93   $ISABELLE_ISAC/Specify
    94     ML_file formalise.sml
    95     ML_file "o-model.sml"
    96     ML_file "i-model.sml"
    97     ML_file "p-model.sml"
    98     ML_file model.sml
    99     ML_file "pre-conditions.sml"
   100     ML_file refine.sml
   101     ML_file "mstools.sml" (*..TODO review*)
   102     ML_file ptyps.sml     (*..TODO review*)
   103     ML_file "test-out.sml"
   104     ML_file "specify-step.sml"
   105     ML_file calchead.sml  (*..TODO review*)
   106     ML_file "input-calchead.sml"
   107     ML_file "step-specify.sml"
   108     ML_file specify.sml
   109 ( ** )    "Specify/Specify"( **)
   110 (*
   111   theory Interpret imports "$ISABELLE_ISAC/Specify/Specify"
   112   $ISABELLE_ISAC/Interpret
   113     ML_file istate.sml
   114     ML_file "sub-problem.sml"
   115     ML_file "thy-read.sml"
   116     ML_file "solve-step.sml"
   117     ML_file "error-pattern.sml"
   118     ML_file derive.sml
   119     ML_file "li-tool.sml"
   120     ML_file "lucas-interpreter.sml"
   121     ML_file "step-solve.sml"
   122 ( ** )    "Interpret/Interpret"( **)
   123 (*
   124   theory MathEngine imports Interpret.Interpret
   125   $ISABELLE_ISAC/MathEngine
   126     ML_file "fetch-tactics.sml"
   127     ML_file solve.sml
   128     ML_file step.sml
   129     ML_file "detail-step.sml"
   130     ML_file "mathengine-stateless.sml"
   131     ML_file messages.sml
   132     ML_file states.sml
   133 ( ** )    "MathEngine/MathEngine"( **)
   134 (*
   135   theory Test_Code imports "$ISABELLE_ISAC/MathEngine/MathEngine"
   136   $ISABELLE_ISAC/Test_Code
   137     ML_file "test-code.sml"
   138 ( ** )    "Test_Code/Test_Code"( **)
   139 (*
   140   theory BridgeLibisabelle imports "$ISABELLE_ISAC/MathEngine/MathEngine"
   141   $ISABELLE_ISAC/BridgeLibisabelle
   142     ML_file "thy-present.sml"
   143     ML_file mathml.sml   
   144     ML_file datatypes.sml
   145     ML_file "pbl-met-hierarchy.sml"
   146     ML_file "thy-hierarchy.sml" 
   147     ML_file "interface-xml.sml"
   148     ML_file interface.sml
   149 ( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
   150 (*
   151     theory Calculation imports Calculation (*preliminary for starting Isabelle/Isac*)
   152     $ISABELLE_ISAC/BridgeJEdit
   153       ML_file parseC.sml
   154       ML_file preliminary.sml
   155   theory BridgeJEdit imports Calculation (*preliminary for starting Isabelle/Isac*)
   156   $ISABELLE_ISAC/BridgeJEdit
   157 ( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
   158 (*
   159     theory Isac imports  "$ISABELLE_ISAC/MathEngine/MathEngine"
   160       ML_file parseC.sml
   161   theory BridgeJEdit imports Isac    
   162 ( **)    "BridgeJEdit/BridgeJEdit"  (*DEactivate after devel.of BridgeJEdit*)
   163 
   164           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
   165 
   166 (*//-----------------------------------------------------------------------------------------\\*)
   167 (*\\-----------------------------------------------------------------------------------------//*)
   168 begin
   169 
   170 text \<open>
   171   show theory dependencies using the graph browser, 
   172   open "browser_info/HOL/Isac/session.graph"
   173   and proceed from the ancestors towards the siblings.
   174 \<close>
   175 
   176 section \<open>check presence of definitions from directories\<close>
   177 
   178 (*declare [[ML_print_depth = 999]]*)
   179 ML \<open>
   180 \<close> ML \<open>
   181 \<close> ML \<open>
   182 \<close>
   183 ML \<open>Eval.adhoc_thm; (*from "ProgLang/evaluate.sml" *)\<close>
   184 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
   185 ML \<open>Input_Descript.for_real_list; (*from "Input_Descript.thy" *)\<close>
   186 ML \<open>Test_Code.me;\<close>
   187 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
   188 ML \<open>prog_expr\<close>
   189                                                          
   190 ML \<open>Prog_Expr.eval_occurs_in\<close>
   191 ML \<open>@{thm last_thmI}\<close>
   192 (** )ML \<open>@{thm Querkraft_Belastung}\<close>( *exception FAIL NONE raised (line 161 of "General/scan.ML")*)
   193 
   194 ML \<open>Check_Unique.on := false;\<close>
   195 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
   196 ML \<open>@{theory "Isac_Knowledge"}\<close>
   197 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   198   ERROR: app_py: not found: ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]*)\<close>
   199 
   200 section \<open>State of approaching Isabelle by Isac\<close>
   201 text \<open>
   202   Mathias Lehnfeld gives the following list in his thesis in section 
   203   4.2.3 Relation to Ongoing Isabelle Development.
   204 \<close>
   205 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   206 text \<open>
   207   REPLACE BY Know_Store... (has been overlooked)
   208     calcelems.sml:val rew_ord' = Unsynchronized.ref ...
   209   KEEP FOR TRACING
   210     rewrite.sml:val Rewrite.trace_on = Unsynchronized.ref false;
   211     rewrite.sml:val depth = Unsynchronized.ref 99999;
   212     rewrite.sml:val lim_rewrite = Unsynchronized.ref 99999;
   213     rewrite.sml:val lim_deriv = Unsynchronized.ref 100;
   214     Interpret/rewtools.sml:val LItool.trace = Unsynchronized.ref false;
   215   KEEP FOR EASIER DEVELOPMENT
   216     calcelems.sml:val check_guhs_unique = Unsynchronized.ref true;
   217   KEEP FOR DEMOS
   218     Knowledge/GCD_Poly_ML.thy:  val trace_div = Unsynchronized.ref true;
   219     Knowledge/GCD_Poly_ML.thy:  val trace_div_invariant = Unsynchronized.ref false;
   220     Knowledge/GCD_Poly_ML.thy:  val trace_Euclid = Unsynchronized.ref true;
   221 \<close>
   222 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   223 subsection \<open>(2) Make Isac’s programming language usable\<close>
   224 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   225 text \<open>
   226   In 2002 isac already strived for floating point numbers. Since that time
   227   isac represents numerals as "Free", see below (*1*). These experiments are
   228   unsatisfactory with respect to logical soundness.
   229   Since Isabelle now has started to care about floating point numbers, it is high 
   230   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   231   are different from Isabelle2011, see "$ISABELLE_ISAC_TEST/Tools/isac/ProgLang/termC.sml".
   232   
   233   The transition from "Free" to standard numerals is a task to be scheduled for 
   234   several weeks. The urgency of this task follows from the experience, 
   235   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   236   some of the long identifiers of theorems which would tremendously simplify
   237   building a hierarchy of theorems according to (1.2), see (*2*) below.
   238 \<close>
   239 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   240 
   241 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   242 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   243 
   244 end