src/Tools/isac/Build_Isac.thy
author wneuper <Walther.Neuper@jku.at>
Mon, 12 Sep 2022 18:02:44 +0200
changeset 60551 3b5be6fae2f0
parent 60549 c0a775618258
child 60556 486223010ea8
permissions -rw-r--r--
eliminate KEStore_Elems.get_thes, add_thes 2: get_cas 2: cleanup finished
     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 Calc_Binop imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    46         at $ISABELLE_ISAC/ProgLang
    47         ML_file evaluate.sml
    48       theory ListC imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    49       theory Program imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    50     theory Prog_Expr imports Calc_Binop ListC Program
    51      theory Prog_Tac imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    52       theory Tactical imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    53     theory Auto_Prog imports Prog_Tac Tactical
    54   theory ProgLang imports Prog_Expr Auto_Prog
    55     at $ISABELLE_ISAC/ProgLang
    56 ( ** )    "ProgLang/ProgLang"( **)
    57 (*
    58   theory MathEngBasic imports
    59     "$ISABELLE_ISAC/ProgLang/ProgLang" "$ISABELLE_ISAC/Specify/Input_Descript"
    60     at $ISABELLE_ISAC/MathEngBasic
    61     ML_file thmC.sml
    62     ML_file problem.sml
    63     ML_file method.sml
    64     ML_file "cas-command.sml"
    65   
    66     ML_file rewrite.sml
    67   
    68     ML_file "model-def.sml"
    69     ML_file "istate-def.sml"
    70     ML_file "calc-tree-elem.sml"
    71     ML_file "pre-conds-def.sml"
    72   
    73     ML_file tactic.sml
    74     ML_file applicable.sml
    75   
    76     ML_file position.sml
    77     ML_file "ctree-basic.sml"
    78     ML_file "ctree-access.sml"
    79     ML_file "ctree-navi.sml"
    80     ML_file ctree.sml
    81   
    82     ML_file "state-steps.sml"
    83     ML_file calculation.sml
    84 ( ** )    "MathEngBasic/MathEngBasic"( **)
    85 (*
    86     theory Input_Descript imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
    87     $ISABELLE_ISAC/Specify
    88   theory Specify imports "$ISABELLE_ISAC/ProgLang/ProgLang" Input_Descript
    89   $ISABELLE_ISAC/Specify
    90     ML_file formalise.sml
    91     ML_file "o-model.sml"
    92     ML_file "i-model.sml"
    93     ML_file "p-model.sml"
    94     ML_file model.sml
    95     ML_file "pre-conditions.sml"
    96     ML_file refine.sml
    97     ML_file "mstools.sml" (*..TODO review*)
    98     ML_file ptyps.sml     (*..TODO review*)
    99     ML_file "test-out.sml"
   100     ML_file "specify-step.sml"
   101     ML_file calchead.sml  (*..TODO review*)
   102     ML_file "input-calchead.sml"
   103     ML_file "step-specify.sml"
   104     ML_file specify.sml
   105 ( ** )    "Specify/Specify"( **)
   106 (*
   107   theory Interpret imports "$ISABELLE_ISAC/Specify/Specify"
   108   $ISABELLE_ISAC/Interpret
   109     ML_file istate.sml
   110     ML_file "sub-problem.sml"
   111     ML_file "thy-read.sml"
   112     ML_file "solve-step.sml"
   113     ML_file "error-pattern.sml"
   114     ML_file derive.sml
   115     ML_file "li-tool.sml"
   116     ML_file "lucas-interpreter.sml"
   117     ML_file "step-solve.sml"
   118 ( ** )    "Interpret/Interpret"( **)
   119 (*
   120   theory MathEngine imports Interpret.Interpret
   121   $ISABELLE_ISAC/MathEngine
   122     ML_file "fetch-tactics.sml"
   123     ML_file solve.sml
   124     ML_file step.sml
   125     ML_file "detail-step.sml"
   126     ML_file "mathengine-stateless.sml"
   127     ML_file messages.sml
   128     ML_file states.sml
   129 ( ** )    "MathEngine/MathEngine"( **)
   130 (*
   131   theory Test_Code imports "$ISABELLE_ISAC/MathEngine/MathEngine"
   132   $ISABELLE_ISAC/Test_Code
   133     ML_file "test-code.sml"
   134 ( ** )    "Test_Code/Test_Code"( **)
   135 (*
   136   theory BridgeLibisabelle imports "$ISABELLE_ISAC/MathEngine/MathEngine"
   137   $ISABELLE_ISAC/BridgeLibisabelle
   138     ML_file "thy-present.sml"
   139     ML_file mathml.sml   
   140     ML_file datatypes.sml
   141     ML_file "pbl-met-hierarchy.sml"
   142     ML_file "thy-hierarchy.sml" 
   143     ML_file "interface-xml.sml"
   144     ML_file interface.sml
   145 ( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
   146 (*
   147     theory Calculation imports Calculation (*preliminary for starting Isabelle/Isac*)
   148     $ISABELLE_ISAC/BridgeJEdit
   149       ML_file parseC.sml
   150       ML_file preliminary.sml
   151   theory BridgeJEdit imports Calculation (*preliminary for starting Isabelle/Isac*)
   152   $ISABELLE_ISAC/BridgeJEdit
   153 ( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
   154 (*
   155     theory Isac imports  "$ISABELLE_ISAC/MathEngine/MathEngine"
   156       ML_file parseC.sml
   157   theory BridgeJEdit imports Isac    
   158 ( **)    "BridgeJEdit/BridgeJEdit"  (*DEactivate after devel.of BridgeJEdit*)
   159 
   160           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
   161 
   162 (*//-----------------------------------------------------------------------------------------\\*)
   163 (*\\-----------------------------------------------------------------------------------------//*)
   164 begin
   165 
   166 text \<open>
   167   show theory dependencies using the graph browser, 
   168   open "browser_info/HOL/Isac/session.graph"
   169   and proceed from the ancestors towards the siblings.
   170 \<close>
   171 
   172 section \<open>check presence of definitions from directories\<close>
   173 
   174 (*declare [[ML_print_depth = 999]]*)
   175 ML \<open>
   176 \<close> ML \<open>
   177 \<close> ML \<open>
   178 \<close>
   179 ML \<open>Eval.adhoc_thm; (*from "ProgLang/evaluate.sml" *)\<close>
   180 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
   181 ML \<open>Input_Descript.for_real_list; (*from "Input_Descript.thy" *)\<close>
   182 ML \<open>Test_Code.me;\<close>
   183 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
   184 ML \<open>prog_expr\<close>
   185                                                          
   186 ML \<open>Prog_Expr.eval_occurs_in\<close>
   187 ML \<open>@{thm last_thmI}\<close>
   188 (** )ML \<open>@{thm Querkraft_Belastung}\<close>( *exception FAIL NONE raised (line 161 of "General/scan.ML")*)
   189 
   190 declare [[check_unique = false]]
   191 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
   192 ML \<open>@{theory "Isac_Knowledge"}\<close>
   193 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   194   ERROR: app_py: not found: ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]*)\<close>
   195 
   196 section \<open>State of approaching Isabelle by Isac\<close>
   197 text \<open>
   198   Mathias Lehnfeld gives the following list in his thesis in section 
   199   4.2.3 Relation to Ongoing Isabelle Development.
   200 \<close>
   201 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   202 text \<open>
   203   DONE
   204 \<close>
   205 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   206 subsection \<open>(2) Make Isac’s programming language usable\<close>
   207 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   208 text \<open>
   209   In 2002 isac already strived for floating point numbers. Since that time
   210   isac represents numerals as "Free", see below (*1*). These experiments are
   211   unsatisfactory with respect to logical soundness.
   212   Since Isabelle now has started to care about floating point numbers, it is high 
   213   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   214   are different from Isabelle2011, see "$ISABELLE_ISAC_TEST/Tools/isac/ProgLang/termC.sml".
   215   
   216   The transition from "Free" to standard numerals is a task to be scheduled for 
   217   several weeks. The urgency of this task follows from the experience, 
   218   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   219   some of the long identifiers of theorems which would tremendously simplify
   220   building a hierarchy of theorems according to (1.2), see (*2*) below.
   221 \<close>
   222 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   223 
   224 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   225 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   226 
   227 end