src/Tools/isac/Build_Isac.thy
author wneuper <Walther.Neuper@jku.at>
Tue, 05 Dec 2023 18:15:45 +0100
changeset 60772 ac0317936138
parent 60771 1b072aab8f4e
child 60773 439e23525491
permissions -rw-r--r--
prepare 1: towards I_Model.T_POS in Specify/ (and no I_Model.T)
     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 "model-def.sml"
    63     ML_file problem.sml
    64     ML_file method.sml
    65     ML_file "cas-command.sml"
    66   
    67     ML_file rewrite.sml
    68   
    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 "o-model.sml"
    91     ML_file "pre-conditions.sml" (*uses Model_Def.i_model*)
    92     ML_file "i-model.sml"
    93     ML_file "p-model.sml"
    94     ML_file "m-match.sml"
    95     ML_file refine.sml
    96     ML_file "test-out.sml"
    97     ML_file "specify-step.sml"
    98     ML_file specification.sml
    99     ML_file "cas-command.sml"
   100     ML_file "p-spec.sml"
   101     ML_file specify.sml
   102     ML_file "sub-problem.sml"
   103     ML_file "step-specify.sml"
   104 ( ** )    "Specify/Specify"( **)
   105 (*
   106   theory Interpret imports "$ISABELLE_ISAC/Specify/Specify"
   107   $ISABELLE_ISAC/Interpret
   108     ML_file istate.sml
   109     ML_file "sub-problem.sml"
   110     ML_file "thy-read.sml"
   111     ML_file "li-tool.sml"
   112     ML_file "solve-step.sml"
   113     ML_file "error-pattern.sml"
   114     ML_file derive.sml
   115     ML_file "lucas-interpreter.sml"
   116     ML_file "step-solve.sml"
   117 ( ** )    "Interpret/Interpret"( **)
   118 (*
   119   theory MathEngine imports Interpret.Interpret
   120   $ISABELLE_ISAC/MathEngine
   121     ML_file "fetch-tactics.sml"
   122     ML_file solve.sml
   123     ML_file step.sml
   124     ML_file "detail-step.sml"
   125     ML_file "mathengine-stateless.sml"
   126     ML_file messages.sml
   127     ML_file states.sml
   128 ( ** )    "MathEngine/MathEngine"( **)
   129 (*
   130   theory Test_Code imports "$ISABELLE_ISAC/MathEngine/MathEngine"
   131   $ISABELLE_ISAC/Test_Code
   132     ML_file "test-code.sml"
   133 ( ** )    "Test_Code/Test_Code"( **)
   134 (*
   135   theory BridgeLibisabelle imports "$ISABELLE_ISAC/MathEngine/MathEngine"
   136   $ISABELLE_ISAC/BridgeLibisabelle
   137     ML_file "thy-present.sml"
   138     ML_file mathml.sml   
   139     ML_file datatypes.sml
   140     ML_file "pbl-met-hierarchy.sml"
   141     ML_file "thy-hierarchy.sml" 
   142     ML_file "interface-xml.sml"
   143     ML_file interface.sml
   144 ( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
   145 (*
   146     theory Calculation imports Calculation (*preliminary for starting Isabelle/Isac*)
   147     $ISABELLE_ISAC/BridgeJEdit
   148       ML_file preliminary.sml
   149   theory BridgeJEdit imports Calculation (*preliminary for starting Isabelle/Isac*)
   150   $ISABELLE_ISAC/BridgeJEdit
   151 ( ** )    "BridgeLibisabelle/BridgeLibisabelle"( **)
   152 (*
   153         theory E_Collect imports "$ISABELLE_ISAC/Knowledge/Build_Thydata"
   154       theory Calculation imports E_Collect "$ISABELLE_ISAC/Knowledge/Build_Isac"
   155         ML_file "user-model.sml"
   156         (*ML_file "template.sml"*)
   157         ML_file "preliminary.sml"
   158     theory VSCode_Example imports Calculation
   159   theory BridgeJEdit imports VSCode_Example    
   160 ( **)    "BridgeJEdit/BridgeJEdit"  (*DEactivate after devel.of BridgeJEdit*)
   161 
   162           "Knowledge/Build_Thydata" (*imports Isac.thy etc*)
   163 
   164 (*//-----------------------------------------------------------------------------------------\\*)
   165 (*\\-----------------------------------------------------------------------------------------//*)
   166 begin
   167 ML \<open>
   168 (** )
   169 ( **)
   170 \<close> ML \<open>
   171 @{term Gleichungen}
   172 \<close> ML \<open>
   173 \<close> 
   174 subsection \<open>make Minisubpbl independent from Thy_Info\<close>
   175 text \<open>
   176   We want Isac to become independent from sessions Specify, Interpret and Isac.
   177   This goal will be checked by the tests finally.
   178   As first step we go top down and make Minisubpbl independent form Thy_Info here
   179   and make sure, that the right ctxt is passed throughout the code.
   180   As next step we go bottom up from Thy_Info.get_theory and remove it.
   181   Afterwards $ISABELLE_ISAC_POS will be changed accordingly.
   182 \<close>
   183 (** ) (* evaluated in Test_Isac/_Short *)
   184   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/000-comments.sml"
   185   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/100a-init-rootpbl-Maximum.sml"
   186   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/150a-add-given-Maximum.sml"
   187   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/150-add-given-Equation.sml"
   188   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/200-start-method-NEXT_STEP.sml"
   189   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/200-start-method.sml"
   190   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/250-Rewrite_Set-from-method-NEXT_STEP.sml"
   191   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/250-Rewrite_Set-from-method.sml"
   192   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/300-init-subpbl-NEXT_STEP.sml"
   193   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/300-init-subpbl.sml"
   194   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/400-start-meth-subpbl.sml"
   195   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/450-Rewrite_Set_Inst.sml"
   196   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/470-Check_elementwise-NEXT_STEP.sml"
   197   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/490-nxt-Check_Postcond.sml"
   198   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/500-met-sub-to-root.sml"
   199   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/530-error-Check_Elementwise.sml"
   200   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/600-postcond-NEXT_STEP.sml"
   201   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/600-postcond.sml"
   202   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/700-interSteps.sml"
   203   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/710-interSteps-short.sml"
   204   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/790-complete-NEXT_STEP.sml"
   205   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/790-complete.sml"
   206   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Minisubpbl/800-append-on-Frm.sml"
   207 (**)                                          
   208   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Specify/i-model.sml"
   209 (**)                                          
   210   ML_file "$ISABELLE_ISAC_POS/Tools/isac/Specify/pre-conditions.sml"
   211 
   212   ML_file "$ISABELLE_ISAC_POS/Tools/isac/BridgeJEdit/e-collect.sml"
   213   ML_file "$ISABELLE_ISAC_POS/Tools/isac/BridgeJEdit/user-model.sml"
   214   ML_file "$ISABELLE_ISAC_POS/Tools/isac/BridgeJEdit/template.sml"
   215   ML_file "$ISABELLE_ISAC_POS/Tools/isac/BridgeJEdit/preliminary.sml"
   216   ML_file "$ISABELLE_ISAC_POS/Tools/isac/BridgeJEdit/calculation.sml"
   217   ML_file "$ISABELLE_ISAC_POS/Tools/isac/BridgeJEdit/vscode-example.sml"
   218 ( **)                                          
   219 
   220 (* evaluated in Test_Isac/_Short *)                               
   221 
   222 section \<open>check presence of definitions from directories\<close>
   223 
   224 (*declare [[ML_print_depth = 999]]*)
   225 ML \<open>Eval.adhoc_thm; (*from "ProgLang/evaluate.sml" *)\<close>
   226 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
   227 ML \<open>Input_Descript.for_real_list; (*from "Input_Descript.thy" *)\<close>
   228 ML \<open>(*Test_Code.me;*)\<close>
   229 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
   230 ML \<open>prog_expr\<close>
   231                                                          
   232 ML \<open>Prog_Expr.eval_occurs_in\<close>
   233 ML \<open>@{thm last_thmI}\<close>
   234 (** )ML \<open>@{thm Querkraft_Belastung}\<close>( *exception FAIL NONE raised (line 161 of "General/scan.ML")*)
   235 
   236 declare [[check_unique = false]]
   237 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
   238 ML \<open>@{theory "Isac_Knowledge"}\<close>
   239 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   240   ERROR: app_py: not found: ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]*)\<close>
   241 
   242 section \<open>State of approaching Isabelle by Isac\<close>
   243 text \<open>
   244   Mathias Lehnfeld gives the following list in his thesis in section 
   245   4.2.3 Relation to Ongoing Isabelle Development.
   246 \<close>
   247 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   248 text \<open>
   249   DONE
   250 \<close>
   251 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   252 subsection \<open>(2) Make Isac’s programming language usable\<close>
   253 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   254 text \<open>
   255   In 2002 isac already strived for floating point numbers. Since that time
   256   isac represents numerals as "Free", see below (*1*). These experiments are
   257   unsatisfactory with respect to logical soundness.
   258   Since Isabelle now has started to care about floating point numbers, it is high 
   259   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   260   are different from Isabelle2011, see "$ISABELLE_ISAC_POS/Tools/isac/ProgLang/termC.sml".
   261   
   262   The transition from "Free" to standard numerals is a task to be scheduled for 
   263   several weeks. The urgency of this task follows from the experience, 
   264   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   265   some of the long identifiers of theorems which would tremendously simplify
   266   building a hierarchy of theorems according to (1.2), see (*2*) below.
   267 \<close>
   268 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   269 
   270 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   271 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   272 
   273 ML \<open>
   274 \<close> text \<open>
   275 I_Model.T_POS
   276 \<close> ML \<open>
   277 I_Model.TEST_to_OLD  ;
   278 I_Model.TEST_to_OLD_single  ;
   279 I_Model.feedback_POS_to_OLD  ;
   280 
   281 I_Model.OLD_to_POS  ;
   282 I_Model.OLD_to_POS_single  ;
   283 I_Model.feedback_OLD_to_POS  ;
   284 (*OLD*)
   285 (*---*)
   286 (*NEW*)
   287 \<close> ML \<open>
   288 \<close>
   289 end
   290 
   291 
   292 
   293 
   294 
   295 
   296 
   297 
   298 
   299 
   300 
   301 
   302 
   303 
   304 
   305 
   306