src/Tools/isac/Build_Isac.thy
author wneuper <Walther.Neuper@jku.at>
Wed, 25 Jan 2023 15:52:33 +0100
changeset 60651 b7a2ad3b3d45
parent 60639 b8bb7d8800e8
child 60652 75003e8f96ab
permissions -rw-r--r--
ContextC.build_while_parsing, improves O_Model.init_PIDE
     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 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 "li-tool.sml"
   113     ML_file "solve-step.sml"
   114     ML_file "error-pattern.sml"
   115     ML_file derive.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 ML \<open>
   166 \<close> ML \<open>
   167 \<close> ML \<open>
   168 \<close> ML \<open>
   169 \<close> 
   170 subsection \<open>make Minisubpbl independent from Thy_Info\<close>
   171 text \<open>
   172   We want Isac to become independent from sessions Specify, Interpret and Isac.
   173   This goal will be checked by the tests finally.
   174   As first step we go top down and make Minisubpbl independent form Thy_Info here
   175   and make sure, that the right ctxt is passed throughout the code.
   176   As next step we go bottom up from Thy_Info.get_theory and remove it.
   177   Afterwards $ISABELLE_ISAC_TEST will be changed accordingly.
   178 \<close>
   179   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/000-comments.sml"
   180   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/100-init-rootpbl.sml"
   181 ML \<open>
   182 \<close> ML \<open>
   183 (* Title:  100-init-rootpbl.sml
   184    Author: Walther Neuper 1105
   185    (c) copyright due to lincense terms.
   186 *)
   187 
   188 "----------- Minisubplb/100-init-rootpbl.sml ---------------------";
   189 "----------- Minisubplb/100-init-rootpbl.sml ---------------------";
   190 "----------- Minisubplb/100-init-rootpbl.sml ---------------------";
   191 val (_(*example text*), 
   192   (model as ("Constants [r = (7::real)]" :: "Maximum A" :: "AdditionalValues [u, v]" :: 
   193      "Extremum (A = 2 * u * v - u \<up> 2)" :: 
   194      "SideConditions [((u::real) / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]" :: 
   195      "SideConditions [((u::real) / 2) \<up> 2 + (2 / v) \<up> 2 = r \<up> 2]" :: 
   196      "SideConditions [(u::real) / 2 = r * sin \<alpha>, 2 / v = r * cos \<alpha>]" :: 
   197      "FunctionVariable a" :: "FunctionVariable b" :: "FunctionVariable \<alpha>" :: 
   198      "Domain {0 <..< r}" :: "Domain {0 <..< r}" :: "Domain {0 <..< \<pi> / 2}" ::
   199      "ErrorBound (\<epsilon> = (0::real))" :: []), 
   200    refs as ("Diff_App", 
   201      ["univariate_calculus", "Optimisation"],
   202      ["Optimisation", "by_univariate_calculus"])))
   203   = Store.get (Know_Store.get_expls @{theory}) ["Diff_App-No.123a"] ["Diff_App-No.123a"];
   204 
   205  Test_Code.init_calc @{context} [(model, refs)];
   206 "~~~~~ fun init_calc , args:"; val (ctxt, [(model, refs as (thy_id, _, _))])
   207   = (@{context}, [(model, refs)]);
   208 (*old* )val ctxt = thy_id |> ThyC.get_theory_PIDE ctxt |> Proof_Context.init_global
   209 ( *new*)val thy = thy_id |> ThyC.get_theory_PIDE ctxt  (*^^^^^^^^^^^^^^^^^^^^^^^^^ misleading*)
   210 
   211 (**)val return_nxt_specify_init_calc_PIDE =(**)
   212 Step_Specify.nxt_specify_init_calc_PIDE thy (model, refs);
   213 "~~~~~ fun nxt_specify_init_calc , args:"; val (ctxt, (model, (dI, pI, mI))) = (ctxt, (model, refs));
   214 
   215 Step_Specify.initialise_PIDE thy (model, refs);
   216 "~~~~~ fun initialise , args:"; val (ctxt, (fmz, (dI, pI, mI))) = (ctxt, (model, refs));
   217 (*old* )val thy = ThyC.get_theory_PIDE ctxt dI( *old*)
   218 	    val ctxt = Proof_Context.init_global thy (* ctxt restricted to Formalise.refs *)
   219 	    val (pI, (pors, pctxt), mI) = 
   220         if mI = ["no_met"] 
   221         then raise error "else not covered by test"
   222         else
   223           let
   224 (*old* )     val pors = Problem.from_store ctxt pI |> #model |> O_Model.init thy fmz; (*..TermC.parseNEW'*)
   225 (*old*)     val pctxt = ContextC.initialise' thy fmz;                (*..DUPLICATE ermC.parseNEW'*)
   226 (*old*)   in (pI, (pors, pctxt), mI) end;
   227 ( *new*)     val (pors, pctxt) = Problem.from_store ctxt pI |> #model |> O_Model.init_PIDE thy fmz;
   228           in (pI, (pors, pctxt), mI) end;
   229 
   230 (*/------------------- check result of O_Model.init_PIDE ------------------------------------\*)
   231 (*+*)val Free ("r", Type ("Real.real", [])) = Syntax.read_term pctxt "r";
   232 (*+*)val Free ("u", Type ("Real.real", [])) = Syntax.read_term pctxt "u";
   233 (*-------------------- stop step into Step_Specify.nxt_specify_init_calc ---------------------*)
   234 
   235 val ((pt, p), tacis) = return_nxt_specify_init_calc_PIDE(* kept for continuation *);
   236 "~~~~~ continue fun init_calc , args:"; val ((pt, p), tacis)
   237   = ((pt, p), tacis);
   238 	  val tac = case tacis of [] => Tactic.Empty_Tac | _ => (#1 o hd) tacis;
   239 
   240  Test_Code.TESTg_form ctxt (pt,p);
   241 "~~~~~ fun TESTg_form , args:"; val (ctxt, ptp) = (ctxt, (pt,p));
   242     val (form, _, _) =  ME_Misc.pt_extract ctxt ptp;
   243 val Ctree.ModSpec (_, p_, _, gfr, where_, _) =
   244     (*case*) form (*of*);
   245 val Pos.Pbl =
   246       (*case*) p_ (*of*);
   247     Test_Out.Problem [];
   248     Test_Out.MethodC [];
   249 
   250 (*val return =*)
   251   Test_Out.PpcKF (Test_Out.Problem [], 
   252    P_Model.from (Proof_Context.theory_of ctxt) gfr where_);
   253 "~~~~~ fun from , args:"; val (thy, itms, where_) = ((Proof_Context.theory_of ctxt), gfr, where_);
   254 
   255 (*//------------------ inserted hidden code ------------------------------------------------\\*)
   256 fun item_from_feedback thy (I_Model.Cor ((d, ts), _)) = P_Model.Correct (UnparseC.term_in_thy thy (Input_Descript.join (d, ts)))
   257   | item_from_feedback _ (I_Model.Syn c) = P_Model.SyntaxE c
   258   | item_from_feedback _ (I_Model.Typ c) = P_Model.TypeE c
   259   | item_from_feedback thy (I_Model.Inc ((d, ts), _)) = P_Model.Incompl (UnparseC.term_in_thy thy  (Input_Descript.join (d, ts)))
   260   | item_from_feedback thy (I_Model.Sup (d, ts)) = P_Model.Superfl (UnparseC.term_in_thy thy  (Input_Descript.join (d, ts)))
   261   | item_from_feedback thy (I_Model.Mis (d, pid)) = P_Model.Missing (UnparseC.term_in_thy thy d ^ " " ^ UnparseC.term_in_thy thy  pid)
   262   | item_from_feedback _ _ = raise ERROR "item_from_feedback: uncovered definition"
   263 fun add_sel_ppc (_: theory) sel {Given = gi, Where = wh, Find = fi, With = wi, Relate = re} x =
   264   case sel of
   265     "#Given" => {Given = gi @ [x], Where = wh, Find = fi, With = wi,Relate = re}
   266   | "#Where" => {Given = gi, Where = wh @ [x], Find = fi, With = wi, Relate = re}
   267   | "#Find"  => {Given = gi, Where = wh, Find = fi @ [x], With = wi, Relate = re}
   268   | "#Relate"=> {Given = gi, Where = wh, Find = fi, With = wi, Relate = re @ [x]}
   269   | "#undef" => {Given = gi @ [x], Where = wh, Find = fi, With = wi, Relate = re} (*ori2itmSup*)
   270   | _  => raise ERROR ("add_sel_ppc tried to select by \"" ^ sel ^ "\"")
   271 fun add_where {Given = gi, Where = _, Find = fi, With = wi, Relate = re} wh =
   272   {Given = gi, Where = wh, Find= fi ,With = wi, Relate = re}
   273 fun boolterm2item ctxt (true, term) = P_Model.Correct (UnparseC.term_in_ctxt ctxt term)
   274   | boolterm2item ctxt(false, term) = P_Model.False (UnparseC.term_in_ctxt ctxt term);
   275 (*\\------------------ inserted hidden code ------------------------------------------------//*)
   276 
   277     fun coll model [] = model
   278       | coll model ((_, _, _, field, itm_)::itms) =
   279         coll (add_sel_ppc thy field model (item_from_feedback thy itm_)) itms;
   280     val gfr = coll P_Model.empty itms;
   281 
   282 (*val return =*)
   283            add_where gfr (map (boolterm2item (Proof_Context.init_global thy)) where_);
   284 "~~~~~ fun add_where , args:"; val ({Given = gi, Where = _, Find = fi, With = wi, Relate = re}, wh)
   285   = (gfr, (map 
   286           (boolterm2item ctxt) where_));
   287 "~~~~~ fun boolterm2item , args:"; val () = ();
   288 (*\\------------------ step into nxt_specify_init_calc -------------------------------------//*)
   289 \<close> ML \<open>
   290 \<close> ML \<open>
   291 \<close> 
   292   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/150a-add-given-Maximum.sml"
   293   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/150-add-given-Equation.sml"
   294   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/200-start-method-NEXT_STEP.sml"
   295   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/200-start-method.sml"
   296   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/250-Rewrite_Set-from-method-NEXT_STEP.sml"
   297   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/250-Rewrite_Set-from-method.sml"
   298   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/300-init-subpbl-NEXT_STEP.sml"
   299   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/300-init-subpbl.sml"
   300   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/400-start-meth-subpbl.sml"
   301   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/450-Rewrite_Set_Inst.sml"
   302   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/470-Check_elementwise-NEXT_STEP.sml"
   303   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/490-nxt-Check_Postcond.sml"
   304   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/500-met-sub-to-root.sml"
   305   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/530-error-Check_Elementwise.sml"
   306   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/600-postcond-NEXT_STEP.sml"
   307   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/600-postcond.sml"
   308   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/700-interSteps.sml"
   309   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/710-interSteps-short.sml"
   310   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/790-complete-NEXT_STEP.sml"
   311   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/790-complete.sml"
   312   ML_file "$ISABELLE_ISAC_TEST/Tools/isac/Minisubpbl/800-append-on-Frm.sml"
   313 
   314 
   315 section \<open>check presence of definitions from directories\<close>
   316 
   317 (*declare [[ML_print_depth = 999]]*)
   318 ML \<open>Eval.adhoc_thm; (*from "ProgLang/evaluate.sml" *)\<close>
   319 ML \<open>Rewrite.rewrite_; (*from "ProgLang/rewrite.sml" *)\<close>
   320 ML \<open>Input_Descript.for_real_list; (*from "Input_Descript.thy" *)\<close>
   321 ML \<open>(*Test_Code.me;*)\<close>
   322 text \<open>contextthyOK2xml; (*"xmlsrc/interface-xml.sml"*)\<close>
   323 ML \<open>prog_expr\<close>
   324                                                          
   325 ML \<open>Prog_Expr.eval_occurs_in\<close>
   326 ML \<open>@{thm last_thmI}\<close>
   327 (** )ML \<open>@{thm Querkraft_Belastung}\<close>( *exception FAIL NONE raised (line 161 of "General/scan.ML")*)
   328 
   329 declare [[check_unique = false]]
   330 ML \<open>writeln "**** isac kernel = math-engine + Knowledge complete ******"\<close>
   331 ML \<open>@{theory "Isac_Knowledge"}\<close>
   332 ML \<open>(*get_the ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]
   333   ERROR: app_py: not found: ["IsacKnowledge", "Diff", "Theorems", "diff_sin_chain"]*)\<close>
   334 
   335 section \<open>State of approaching Isabelle by Isac\<close>
   336 text \<open>
   337   Mathias Lehnfeld gives the following list in his thesis in section 
   338   4.2.3 Relation to Ongoing Isabelle Development.
   339 \<close>
   340 subsection \<open>(0) Survey on remaining Unsynchronized.ref\<close>
   341 text \<open>
   342   DONE
   343 \<close>
   344 subsection \<open>(1) Exploit parallelism for concurrent session management\<close>
   345 subsection \<open>(2) Make Isac’s programming language usable\<close>
   346 subsection \<open>(3) Adopt Isabelle’s numeral computation for Isac\<close>
   347 text \<open>
   348   In 2002 isac already strived for floating point numbers. Since that time
   349   isac represents numerals as "Free", see below (*1*). These experiments are
   350   unsatisfactory with respect to logical soundness.
   351   Since Isabelle now has started to care about floating point numbers, it is high 
   352   time to adopt these together with the other numerals. Isabelle2012/13's numerals
   353   are different from Isabelle2011, see "$ISABELLE_ISAC_TEST/Tools/isac/ProgLang/termC.sml".
   354   
   355   The transition from "Free" to standard numerals is a task to be scheduled for 
   356   several weeks. The urgency of this task follows from the experience, 
   357   that (1.2) for "thehier" is very hard, because "num_str" seems to destroy 
   358   some of the long identifiers of theorems which would tremendously simplify
   359   building a hierarchy of theorems according to (1.2), see (*2*) below.
   360 \<close>
   361 ML \<open>(*1*) Free ("123.456", HOLogic.realT)\<close>
   362 
   363 subsection \<open>(4) Improve the efficiency of Isac’s rewrite-engine\<close>
   364 subsection \<open>(5) Adopt Isabelle/jEdit for Isac\<close>
   365 
   366 end